aboutsummaryrefslogtreecommitdiff
path: root/src/Compilers
diff options
context:
space:
mode:
Diffstat (limited to 'src/Compilers')
-rw-r--r--src/Compilers/CommonSubexpressionElimination.v259
-rw-r--r--src/Compilers/CommonSubexpressionEliminationDenote.v120
-rw-r--r--src/Compilers/CommonSubexpressionEliminationInterp.v218
-rw-r--r--src/Compilers/CommonSubexpressionEliminationProperties.v189
-rw-r--r--src/Compilers/CommonSubexpressionEliminationWf.v223
-rw-r--r--src/Compilers/Conversion.v110
-rw-r--r--src/Compilers/CountLets.v66
-rw-r--r--src/Compilers/Equality.v90
-rw-r--r--src/Compilers/Eta.v75
-rw-r--r--src/Compilers/EtaInterp.v118
-rw-r--r--src/Compilers/EtaWf.v122
-rw-r--r--src/Compilers/ExprInversion.v359
-rw-r--r--src/Compilers/FilterLive.v70
-rw-r--r--src/Compilers/FoldTypes.v45
-rw-r--r--src/Compilers/GeneralizeVar.v39
-rw-r--r--src/Compilers/GeneralizeVarInterp.v79
-rw-r--r--src/Compilers/GeneralizeVarWf.v77
-rw-r--r--src/Compilers/InSet/Syntax.v84
-rw-r--r--src/Compilers/InSet/Typeify.v58
-rw-r--r--src/Compilers/InSet/TypeifyInterp.v114
-rw-r--r--src/Compilers/Inline.v104
-rw-r--r--src/Compilers/InlineConstAndOp.v149
-rw-r--r--src/Compilers/InlineConstAndOpByRewrite.v93
-rw-r--r--src/Compilers/InlineConstAndOpByRewriteInterp.v136
-rw-r--r--src/Compilers/InlineConstAndOpByRewriteWf.v171
-rw-r--r--src/Compilers/InlineConstAndOpInterp.v161
-rw-r--r--src/Compilers/InlineConstAndOpWf.v219
-rw-r--r--src/Compilers/InlineInterp.v137
-rw-r--r--src/Compilers/InlineWf.v231
-rw-r--r--src/Compilers/InputSyntax.v252
-rw-r--r--src/Compilers/InterpByIso.v33
-rw-r--r--src/Compilers/InterpByIsoProofs.v117
-rw-r--r--src/Compilers/InterpProofs.v66
-rw-r--r--src/Compilers/InterpRewriting.v197
-rw-r--r--src/Compilers/InterpSideConditions.v40
-rw-r--r--src/Compilers/InterpWf.v80
-rw-r--r--src/Compilers/InterpWfRel.v94
-rw-r--r--src/Compilers/Intros.v94
-rw-r--r--src/Compilers/Linearize.v97
-rw-r--r--src/Compilers/LinearizeInterp.v136
-rw-r--r--src/Compilers/LinearizeWf.v179
-rw-r--r--src/Compilers/Map.v30
-rw-r--r--src/Compilers/MapBaseType.v116
-rw-r--r--src/Compilers/MapBaseTypeWf.v117
-rw-r--r--src/Compilers/MapCastByDeBruijn.v62
-rw-r--r--src/Compilers/MapCastByDeBruijnInterp.v129
-rw-r--r--src/Compilers/MapCastByDeBruijnWf.v108
-rw-r--r--src/Compilers/MultiSizeTest.v277
-rw-r--r--src/Compilers/Named/AListContext.v143
-rw-r--r--src/Compilers/Named/Compile.v59
-rw-r--r--src/Compilers/Named/CompileInterp.v207
-rw-r--r--src/Compilers/Named/CompileInterpSideConditions.v253
-rw-r--r--src/Compilers/Named/CompileProperties.v74
-rw-r--r--src/Compilers/Named/CompileWf.v254
-rw-r--r--src/Compilers/Named/Context.v65
-rw-r--r--src/Compilers/Named/ContextDefinitions.v67
-rw-r--r--src/Compilers/Named/ContextOn.v24
-rw-r--r--src/Compilers/Named/ContextProperties.v181
-rw-r--r--src/Compilers/Named/ContextProperties/NameUtil.v161
-rw-r--r--src/Compilers/Named/ContextProperties/Proper.v142
-rw-r--r--src/Compilers/Named/ContextProperties/SmartMap.v299
-rw-r--r--src/Compilers/Named/ContextProperties/Tactics.v101
-rw-r--r--src/Compilers/Named/CountLets.v49
-rw-r--r--src/Compilers/Named/DeadCodeElimination.v54
-rw-r--r--src/Compilers/Named/DeadCodeEliminationInterp.v67
-rw-r--r--src/Compilers/Named/EstablishLiveness.v105
-rw-r--r--src/Compilers/Named/ExprInversion.v21
-rw-r--r--src/Compilers/Named/FMapContext.v70
-rw-r--r--src/Compilers/Named/GetNames.v38
-rw-r--r--src/Compilers/Named/IdContext.v27
-rw-r--r--src/Compilers/Named/InterpSideConditions.v54
-rw-r--r--src/Compilers/Named/InterpSideConditionsInterp.v45
-rw-r--r--src/Compilers/Named/InterpretToPHOAS.v65
-rw-r--r--src/Compilers/Named/InterpretToPHOASInterp.v89
-rw-r--r--src/Compilers/Named/InterpretToPHOASWf.v139
-rw-r--r--src/Compilers/Named/MapCast.v72
-rw-r--r--src/Compilers/Named/MapCastInterp.v290
-rw-r--r--src/Compilers/Named/MapCastWf.v285
-rw-r--r--src/Compilers/Named/MapType.v59
-rw-r--r--src/Compilers/Named/NameUtil.v56
-rw-r--r--src/Compilers/Named/NameUtilProperties.v223
-rw-r--r--src/Compilers/Named/PositiveContext.v8
-rw-r--r--src/Compilers/Named/PositiveContext/Defaults.v29
-rw-r--r--src/Compilers/Named/PositiveContext/DefaultsProperties.v38
-rw-r--r--src/Compilers/Named/RegisterAssign.v89
-rw-r--r--src/Compilers/Named/RegisterAssignInterp.v239
-rw-r--r--src/Compilers/Named/SmartMap.v20
-rw-r--r--src/Compilers/Named/Syntax.v92
-rw-r--r--src/Compilers/Named/WeakListContext.v10
-rw-r--r--src/Compilers/Named/Wf.v56
-rw-r--r--src/Compilers/Named/WfFromUnit.v82
-rw-r--r--src/Compilers/Named/WfInterp.v41
-rw-r--r--src/Compilers/Reify.v591
-rw-r--r--src/Compilers/Relations.v368
-rw-r--r--src/Compilers/RenameBinders.v78
-rw-r--r--src/Compilers/Rewriter.v39
-rw-r--r--src/Compilers/RewriterInterp.v50
-rw-r--r--src/Compilers/RewriterWf.v61
-rw-r--r--src/Compilers/SmartMap.v453
-rw-r--r--src/Compilers/StripExpr.v35
-rw-r--r--src/Compilers/Syntax.v154
-rw-r--r--src/Compilers/TestCase.v276
-rw-r--r--src/Compilers/Tuple.v88
-rw-r--r--src/Compilers/TypeInversion.v193
-rw-r--r--src/Compilers/TypeUtil.v35
-rw-r--r--src/Compilers/Wf.v68
-rw-r--r--src/Compilers/WfInversion.v208
-rw-r--r--src/Compilers/WfProofs.v473
-rw-r--r--src/Compilers/WfReflective.v280
-rw-r--r--src/Compilers/WfReflectiveGen.v334
-rw-r--r--src/Compilers/Z/ArithmeticSimplifier.v611
-rw-r--r--src/Compilers/Z/ArithmeticSimplifierInterp.v247
-rw-r--r--src/Compilers/Z/ArithmeticSimplifierUtil.v79
-rw-r--r--src/Compilers/Z/ArithmeticSimplifierWf.v218
-rw-r--r--src/Compilers/Z/BinaryNotationConstants.v3901
-rw-r--r--src/Compilers/Z/Bounds/Interpretation.v289
-rw-r--r--src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v214
-rw-r--r--src/Compilers/Z/Bounds/InterpretationLemmas/PullCast.v192
-rw-r--r--src/Compilers/Z/Bounds/InterpretationLemmas/Tactics.v197
-rw-r--r--src/Compilers/Z/Bounds/MapCastByDeBruijn.v24
-rw-r--r--src/Compilers/Z/Bounds/MapCastByDeBruijnInterp.v48
-rw-r--r--src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v42
-rw-r--r--src/Compilers/Z/Bounds/Pipeline.v56
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/Definition.v285
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/Glue.v550
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/OutputType.v52
-rw-r--r--src/Compilers/Z/Bounds/Pipeline/ReflectiveTactics.v362
-rw-r--r--src/Compilers/Z/Bounds/Relax.v135
-rw-r--r--src/Compilers/Z/Bounds/RoundUpLemmas.v34
-rw-r--r--src/Compilers/Z/CNotations.v12696
-rw-r--r--src/Compilers/Z/CommonSubexpressionElimination.v230
-rw-r--r--src/Compilers/Z/CommonSubexpressionEliminationInterp.v23
-rw-r--r--src/Compilers/Z/CommonSubexpressionEliminationWf.v29
-rw-r--r--src/Compilers/Z/FoldTypes.v17
-rw-r--r--src/Compilers/Z/GeneralizeVar.v15
-rw-r--r--src/Compilers/Z/GeneralizeVarInterp.v23
-rw-r--r--src/Compilers/Z/GeneralizeVarWf.v30
-rw-r--r--src/Compilers/Z/HexNotationConstants.v7043
-rw-r--r--src/Compilers/Z/Inline.v10
-rw-r--r--src/Compilers/Z/InlineConstAndOp.v11
-rw-r--r--src/Compilers/Z/InlineConstAndOpByRewrite.v13
-rw-r--r--src/Compilers/Z/InlineConstAndOpByRewriteInterp.v12
-rw-r--r--src/Compilers/Z/InlineConstAndOpByRewriteWf.v13
-rw-r--r--src/Compilers/Z/InlineConstAndOpInterp.v11
-rw-r--r--src/Compilers/Z/InlineConstAndOpWf.v11
-rw-r--r--src/Compilers/Z/InlineInterp.v15
-rw-r--r--src/Compilers/Z/InlineWf.v15
-rw-r--r--src/Compilers/Z/InterpSideConditions.v9
-rw-r--r--src/Compilers/Z/JavaNotations.v882
-rw-r--r--src/Compilers/Z/MapCastByDeBruijn.v28
-rw-r--r--src/Compilers/Z/MapCastByDeBruijnInterp.v56
-rw-r--r--src/Compilers/Z/MapCastByDeBruijnWf.v56
-rw-r--r--src/Compilers/Z/Named/DeadCodeElimination.v14
-rw-r--r--src/Compilers/Z/Named/DeadCodeEliminationInterp.v50
-rw-r--r--src/Compilers/Z/Named/RewriteAddToAdc.v147
-rw-r--r--src/Compilers/Z/Named/RewriteAddToAdcInterp.v449
-rw-r--r--src/Compilers/Z/OpInversion.v29
-rw-r--r--src/Compilers/Z/Reify.v108
-rw-r--r--src/Compilers/Z/RewriteAddToAdc.v58
-rw-r--r--src/Compilers/Z/RewriteAddToAdcInterp.v71
-rw-r--r--src/Compilers/Z/RewriteAddToAdcWf.v39
-rw-r--r--src/Compilers/Z/Syntax.v103
-rw-r--r--src/Compilers/Z/Syntax/Equality.v215
-rw-r--r--src/Compilers/Z/Syntax/Util.v317
-rw-r--r--src/Compilers/Z/TypeInversion.v43
-rw-r--r--src/Compilers/ZExtended/InlineConstAndOp.v10
-rw-r--r--src/Compilers/ZExtended/InlineConstAndOpByRewrite.v12
-rw-r--r--src/Compilers/ZExtended/InlineConstAndOpByRewriteInterp.v15
-rw-r--r--src/Compilers/ZExtended/InlineConstAndOpByRewriteWf.v13
-rw-r--r--src/Compilers/ZExtended/InlineConstAndOpInterp.v14
-rw-r--r--src/Compilers/ZExtended/InlineConstAndOpWf.v11
-rw-r--r--src/Compilers/ZExtended/MapBaseType.v22
-rw-r--r--src/Compilers/ZExtended/Syntax.v89
-rw-r--r--src/Compilers/ZExtended/Syntax/Util.v81
174 files changed, 0 insertions, 45160 deletions
diff --git a/src/Compilers/CommonSubexpressionElimination.v b/src/Compilers/CommonSubexpressionElimination.v
deleted file mode 100644
index c5ffeb088..000000000
--- a/src/Compilers/CommonSubexpressionElimination.v
+++ /dev/null
@@ -1,259 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Coq.Lists.List.
-Require Import Coq.FSets.FMapInterface.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Equality.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.AListContext.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.Decidable.
-
-Local Open Scope list_scope.
-
-Inductive symbolic_expr {base_type_code op_code} : Type :=
-| STT
-| SVar (n : nat)
-| SOp (argT : flat_type base_type_code) (op : op_code) (args : symbolic_expr)
-| SPair (x y : symbolic_expr)
-| SFst (A B : flat_type base_type_code) (x : symbolic_expr)
-| SSnd (A B : flat_type base_type_code) (x : symbolic_expr)
-| SInvalid.
-Scheme Equality for symbolic_expr.
-
-Arguments symbolic_expr : clear implicits.
-
-Global Instance symbolic_expr_dec {base_type_code op_code}
- `{DecidableRel (@eq base_type_code), DecidableRel (@eq op_code)}
- : DecidableRel (@eq (symbolic_expr base_type_code op_code)).
-Proof.
- unfold Decidable in *; intros; repeat decide equality.
-Defined.
-
-Ltac inversion_symbolic_expr_step :=
- match goal with
- | [ H : SVar _ = SVar _ |- _ ] => inversion H; clear H
- | [ H : SOp _ _ _ = SOp _ _ _ |- _ ] => inversion H; clear H
- | [ H : SPair _ _ = SPair _ _ |- _ ] => inversion H; clear H
- | [ H : SFst _ _ _ = SFst _ _ _ |- _ ] => inversion H; clear H
- | [ H : SSnd _ _ _ = SSnd _ _ _ |- _ ] => inversion H; clear H
- end.
-Ltac inversion_symbolic_expr := repeat inversion_symbolic_expr_step.
-
-Local Open Scope ctype_scope.
-Section symbolic.
- (** Holds decidably-equal versions of raw expressions, for lookup. *)
- Context (base_type_code : Type)
- (op_code : Type)
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (op_code_beq : op_code -> op_code -> bool)
- (base_type_code_bl : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (op_code_bl : forall x y, op_code_beq x y = true -> x = y)
- (op_code_lb : forall x y, x = y -> op_code_beq x y = true)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (symbolize_op : forall s d, op s d -> op_code)
- (op_code_leb : op_code -> op_code -> bool)
- (base_type_leb : base_type_code -> base_type_code -> bool).
- Local Notation symbolic_expr := (symbolic_expr base_type_code op_code).
- Context (normalize_symbolic_op_arguments : op_code -> symbolic_expr -> symbolic_expr)
- (inline_symbolic_expr_in_lookup : bool).
-
- Local Notation symbolic_expr_beq := (@symbolic_expr_beq base_type_code op_code base_type_code_beq op_code_beq).
- Local Notation symbolic_expr_lb := (@internal_symbolic_expr_dec_lb base_type_code op_code base_type_code_beq op_code_beq base_type_code_lb op_code_lb).
- Local Notation symbolic_expr_bl := (@internal_symbolic_expr_dec_bl base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op_code_bl).
-
- Local Notation flat_type_beq := (@flat_type_beq base_type_code base_type_code_beq).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Definition SymbolicExprContext {var : flat_type -> Type}
- : Context symbolic_expr var
- := @AListContext symbolic_expr symbolic_expr_beq _ var flat_type_beq (@flat_type_dec_bl _ _ base_type_code_bl).
-
- Local Instance SymbolicExprContextOk {var} : ContextOk (@SymbolicExprContext var)
- := @AListContextOk
- symbolic_expr symbolic_expr_beq symbolic_expr_bl symbolic_expr_lb
- _ _ _ _ (@flat_type_dec_lb _ _ base_type_code_lb).
-
- Fixpoint push_pair_symbolic_expr {t : flat_type} (s : symbolic_expr)
- : interp_flat_type (fun _ => symbolic_expr) t
- := match t with
- | Unit => tt
- | Tbase T => s
- | Prod A B
- => (@push_pair_symbolic_expr A (SFst A B s), @push_pair_symbolic_expr B (SSnd A B s))
- end.
-
- Section with_var0.
- Context {var : base_type_code -> Type}.
-
- Fixpoint prepend_prefix {t} (e : @exprf var t) (ls : list (sigT (fun t : flat_type => @exprf var t)))
- : @exprf var t
- := match ls with
- | nil => e
- | x :: xs => LetIn (projT2 x) (fun _ => @prepend_prefix _ e xs)
- end.
- End with_var0.
-
- Local Notation op_code_eqb argT1 argT2 op1 op2
- := (flat_type_beq _ base_type_code_beq argT1 argT2 && op_code_beq op1 op2).
- Local Notation ltb_of_leb leb eqb
- := (leb && negb eqb).
- Local Notation leb_pair leb1 eqb1 leb2
- := (ltb_of_leb leb1 eqb1 || (eqb1 && leb2)).
- Local Notation op_code_ltb op1 op2
- := (ltb_of_leb (op_code_leb op1 op2)
- (op_code_beq op1 op2)).
-
- Fixpoint flat_type_leb (x y : flat_type) : bool
- := match x, y with
- | Unit, _ => true
- | _, Unit => false
- | Tbase x, Tbase y => base_type_leb x y
- | Tbase _, _ => true
- | _, Tbase _ => false
- | Prod A1 B1, Prod A2 B2
- => leb_pair (flat_type_leb A1 A2)
- (flat_type_beq A1 A2)
- (flat_type_leb B1 B2)
- end.
-
- Fixpoint symbolic_expr_leb (x y : symbolic_expr) : bool
- := match x, y with
- | STT, _ => true
- | _, STT => false
- | SVar n1, SVar n2 => Nat.leb n1 n2
- | SVar _, _ => true
- | _, SVar _ => false
- | SOp argT1 op1 args1, SOp argT2 op2 args2
- => leb_pair (leb_pair (op_code_leb op1 op2)
- (op_code_beq op1 op2)
- (flat_type_leb argT1 argT2))
- (op_code_beq op1 op2 && flat_type_beq argT1 argT2)
- (symbolic_expr_leb args1 args2)
- | SOp _ _ _, _ => true
- | _, SOp _ _ _ => false
- | SPair x1 y1, SPair x2 y2
- => leb_pair (symbolic_expr_leb x1 x2)
- (symbolic_expr_beq x1 x2)
- (symbolic_expr_leb y1 y2)
- | SPair _ _, _ => true
- | _, SPair _ _ => false
- | SFst A B x, SFst A' B' y
- => leb_pair (flat_type_leb (Prod A B) (Prod A' B'))
- (flat_type_beq (Prod A B) (Prod A' B'))
- (symbolic_expr_leb x y)
- | SFst _ _ _, _ => true
- | _, SFst _ _ _ => false
- | SSnd A B x, SSnd A' B' y
- => leb_pair (flat_type_leb (Prod A B) (Prod A' B'))
- (flat_type_beq (Prod A B) (Prod A' B'))
- (symbolic_expr_leb x y)
- | SSnd _ _ _, _ => true
- | _, SSnd _ _ _ => false
- | SInvalid, _ => true
- end.
- Definition symbolic_expr_ltb x y
- := ltb_of_leb (symbolic_expr_leb x y) (symbolic_expr_beq x y).
-
- Fixpoint symbolic_expr_normalize (x : symbolic_expr) : symbolic_expr
- := match x with
- | STT => STT
- | SVar n => SVar n
- | SOp argT op args
- => SOp argT op (normalize_symbolic_op_arguments op args)
- | SPair x y
- => SPair (symbolic_expr_normalize x) (symbolic_expr_normalize y)
- | SFst A B x => SFst A B (symbolic_expr_normalize x)
- | SSnd A B x => SFst A B (symbolic_expr_normalize x)
- | SInvalid => SInvalid
- end.
-
- Section with_var.
- Context {var : base_type_code -> Type}.
-
- Local Notation svar t := (var t * symbolic_expr)%type.
- Local Notation fsvar := (fun t => svar t).
- Local Notation mapping := (@SymbolicExprContext (interp_flat_type var)).
-
- Context (prefix : list (sigT (fun t : flat_type => @exprf fsvar t))).
-
- Definition symbolize_var (xs : mapping) (t : flat_type) : symbolic_expr :=
- SVar (length xs).
-
- Fixpoint symbolize_exprf
- {t} (v : @exprf fsvar t) {struct v}
- : option symbolic_expr
- := match v with
- | TT => Some STT
- | Var _ x => Some (snd x)
- | Op argsT _ op args => option_map
- (fun sargs => SOp argsT (symbolize_op _ _ op) sargs)
- (@symbolize_exprf _ args)
- | LetIn _ ex _ eC => None
- | Pair _ ex _ ey => match @symbolize_exprf _ ex, @symbolize_exprf _ ey with
- | Some sx, Some sy => Some (SPair sx sy)
- | _, _ => None
- end
- end.
-
- Definition norm_symbolize_exprf {t} v : option symbolic_expr
- := option_map symbolic_expr_normalize (@symbolize_exprf t v).
-
- Definition symbolicify_smart_var {t : flat_type}
- (vs : interp_flat_type var t)
- (ss : symbolic_expr)
- : interp_flat_type fsvar t
- := SmartVarfMap2 (fun t v s => (v, s)) vs (push_pair_symbolic_expr ss).
-
- Definition csef_step
- (csef : forall {t} (v : @exprf fsvar t) (xs : mapping), @exprf var t)
- {t} (v : @exprf fsvar t) (xs : mapping)
- : @exprf var t
- := match v in @Syntax.exprf _ _ _ t return exprf t with
- | LetIn tx ex _ eC
- => let sx := norm_symbolize_exprf ex in
- let ex' := @csef _ ex xs in
- let '(sx, sv) := match sx with
- | Some sx => (sx, lookupb xs sx)
- | None => (symbolize_var xs tx, None)
- end in
- let reduced_sx := if inline_symbolic_expr_in_lookup then sx else symbolize_var xs tx in
- match sv with
- | Some v => @csef _ (eC (symbolicify_smart_var v reduced_sx)) (extendb xs reduced_sx v)
- | None
- => LetIn ex' (fun x => let sx' := symbolicify_smart_var x reduced_sx in
- @csef _ (eC sx') (extendb (extendb xs sx x) reduced_sx x))
- end
- | TT => TT
- | Var _ x => Var (fst x)
- | Op _ _ op args => Op op (@csef _ args xs)
- | Pair _ ex _ ey => Pair (@csef _ ex xs) (@csef _ ey xs)
- end.
-
- Fixpoint csef {t} (v : @exprf fsvar t) (xs : mapping)
- := @csef_step (@csef) t v xs.
-
- Definition cse {t} (v : @expr fsvar t) (xs : mapping) : @expr var t
- := match v in @Syntax.expr _ _ _ t return expr t with
- | Abs src dst f
- => let sx := symbolize_var xs src in
- Abs (fun x => let x' := symbolicify_smart_var x sx in
- csef (prepend_prefix (f x') prefix) (extendb xs sx x))
- end.
- End with_var.
-
- Definition CSE {t} (e : Expr t) (prefix : forall var, list (sigT (fun t : flat_type => @exprf var t)))
- : Expr t
- := fun var => cse (prefix _) (e _) empty.
-End symbolic.
-
-Global Arguments csef {_} op_code base_type_code_beq op_code_beq base_type_code_bl {_} symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup {var t} _ _.
-Global Arguments cse {_} op_code base_type_code_beq op_code_beq base_type_code_bl {_} symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup {var} prefix {t} _ _.
-Global Arguments CSE {_} op_code base_type_code_beq op_code_beq base_type_code_bl {_} symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup {t} e prefix var.
diff --git a/src/Compilers/CommonSubexpressionEliminationDenote.v b/src/Compilers/CommonSubexpressionEliminationDenote.v
deleted file mode 100644
index 2218f62eb..000000000
--- a/src/Compilers/CommonSubexpressionEliminationDenote.v
+++ /dev/null
@@ -1,120 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.AListContext.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Equality.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.CommonSubexpressionElimination.
-
-Section symbolic.
- Context (base_type_code : Type)
- (op_code : Type)
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (op_code_beq : op_code -> op_code -> bool)
- (base_type_code_bl : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (op_code_bl : forall x y, op_code_beq x y = true -> x = y)
- (op_code_lb : forall x y, x = y -> op_code_beq x y = true)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (symbolize_op : forall s d, op s d -> op_code)
- (denote_op : forall s d, op_code -> option (op s d))
- (denote_symbolize_op : forall s d opc, denote_op s d (symbolize_op s d opc) = Some opc).
-
- Local Notation symbolic_expr := (symbolic_expr base_type_code op_code).
- Local Notation symbolic_expr_beq := (@symbolic_expr_beq base_type_code op_code base_type_code_beq op_code_beq).
- Local Notation symbolic_expr_lb := (@internal_symbolic_expr_dec_lb base_type_code op_code base_type_code_beq op_code_beq base_type_code_lb op_code_lb).
- Local Notation symbolic_expr_bl := (@internal_symbolic_expr_dec_bl base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op_code_bl).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Local Notation flat_type_beq := (@flat_type_beq base_type_code base_type_code_beq).
- Local Notation flat_type_dec_bl := (@flat_type_dec_bl base_type_code base_type_code_beq base_type_code_bl).
-
- Local Notation symbolicify_smart_var := (@symbolicify_smart_var base_type_code op_code).
- Local Notation symbolize_exprf := (@symbolize_exprf base_type_code op_code op symbolize_op).
- Local Notation csef := (@csef base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op).
- Local Notation cse := (@cse base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op).
- Local Notation CSE := (@CSE base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op).
- Local Notation SymbolicExprContext := (@SymbolicExprContext base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl).
- Local Notation SymbolicExprContextOk := (@SymbolicExprContextOk base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl base_type_code_lb op_code_bl op_code_lb).
- Local Notation prepend_prefix := (@prepend_prefix base_type_code op).
-
- Section with_var.
- Context {interp_base_type : base_type_code -> Type}
- {interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- (m : @SymbolicExprContext (interp_flat_type interp_base_type)).
-
- Local Notation var_cast := (@var_cast _ (interp_flat_type interp_base_type) flat_type_beq flat_type_dec_bl).
- Fixpoint denote_symbolic_expr
- (t : flat_type)
- (se : symbolic_expr)
- : option (interp_flat_type interp_base_type t)
- := match se, t with
- | STT, Unit => Some tt
- | SVar n, t
- => match List.nth_error m (length m - S n) with
- | Some e => @var_cast _ t (projT2 (snd e))
- | None => None
- end
- | SOp argsT op args, _
- => match denote_op argsT t op, @denote_symbolic_expr argsT args with
- | Some opc, Some eargs => Some (interp_op _ _ opc eargs)
- | Some _, None | None, Some _ | None, None => None
- end
- | SPair x y, Prod A B
- => match @denote_symbolic_expr A x, @denote_symbolic_expr B y with
- | Some ex, Some ey => Some (ex, ey)
- | Some _, None | None, Some _ | None, None => None
- end
- | SFst A B x, A'
- => if flat_type_beq A A'
- then option_map (@fst _ _) (@denote_symbolic_expr (Prod A' B) x)
- else None
- | SSnd A B x, B'
- => if flat_type_beq B B'
- then option_map (@snd _ _) (@denote_symbolic_expr (Prod A B') x)
- else None
- | SInvalid, _
- | STT, _
- | SPair _ _, _
- => None
- end.
-
- Fail Lemma denote_symbolize_exprf
- (*(Hm : forall n v, List.nth_error m n = Some v -> denote_symbolic_expr *)
- t e se e'
- : @symbolize_exprf var t e = Some se
- -> denote_symbolic_expr t se = Some e'
- -> e' = e. (* The command has indeed failed with message:
-In environment
-base_type_code : Type
-op_code : Type
-base_type_code_beq : base_type_code -> base_type_code -> bool
-op_code_beq : op_code -> op_code -> bool
-base_type_code_bl : forall x y : base_type_code, base_type_code_beq x y = true -> x = y
-base_type_code_lb : forall x y : base_type_code, x = y -> base_type_code_beq x y = true
-op_code_bl : forall x y : op_code, op_code_beq x y = true -> x = y
-op_code_lb : forall x y : op_code, x = y -> op_code_beq x y = true
-op : flat_type -> flat_type -> Type
-symbolize_op : forall s d : flat_type, op s d -> op_code
-denote_op : forall s d : flat_type, op_code -> option (op s d)
-denote_symbolize_op : forall (s d : flat_type) (opc : op s d), denote_op s d (symbolize_op s d opc) = Some opc
-var : base_type_code -> Type
-m : SymbolicExprContext
-t : flat_type
-e : exprf t
-se : symbolic_expr
-e' : exprf t
-The term "e" has type
- "@Syntax.exprf base_type_code op
- (fun t : base_type_code => prod (var t) (CommonSubexpressionElimination.symbolic_expr base_type_code op_code)) t"
-while it is expected to have type "@Syntax.exprf base_type_code op var t".
- *)
- End with_var.
-End symbolic.
diff --git a/src/Compilers/CommonSubexpressionEliminationInterp.v b/src/Compilers/CommonSubexpressionEliminationInterp.v
deleted file mode 100644
index 2152b5466..000000000
--- a/src/Compilers/CommonSubexpressionEliminationInterp.v
+++ /dev/null
@@ -1,218 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.AListContext.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Equality.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.CommonSubexpressionElimination.
-Require Import Crypto.Compilers.CommonSubexpressionEliminationDenote.
-Require Import Crypto.Compilers.CommonSubexpressionEliminationWf.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Decidable.
-
-Section symbolic.
- Context (base_type_code : Type)
- (op_code : Type)
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (op_code_beq : op_code -> op_code -> bool)
- (base_type_code_bl : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (op_code_bl : forall x y, op_code_beq x y = true -> x = y)
- (op_code_lb : forall x y, x = y -> op_code_beq x y = true)
- (interp_base_type : base_type_code -> Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (interp_op : forall s d (opc : op s d), interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d)
- (symbolize_op : forall s d, op s d -> op_code)
- (denote_op : forall s d, op_code -> option (op s d)).
- Local Notation symbolic_expr := (symbolic_expr base_type_code op_code).
- Context (normalize_symbolic_op_arguments : op_code -> symbolic_expr -> symbolic_expr)
- (inline_symbolic_expr_in_lookup : bool).
-
- Local Notation symbolic_expr_beq := (@symbolic_expr_beq base_type_code op_code base_type_code_beq op_code_beq).
- Local Notation symbolic_expr_lb := (@internal_symbolic_expr_dec_lb base_type_code op_code base_type_code_beq op_code_beq base_type_code_lb op_code_lb).
- Local Notation symbolic_expr_bl := (@internal_symbolic_expr_dec_bl base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op_code_bl).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type_gen := interp_flat_type.
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Local Notation symbolicify_smart_var := (@symbolicify_smart_var base_type_code op_code).
- Local Notation symbolize_exprf := (@symbolize_exprf base_type_code op_code op symbolize_op).
- Local Notation csef := (@csef base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup).
- Local Notation cse := (@cse base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup).
- Local Notation CSE := (@CSE base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup).
- Local Notation SymbolicExprContext := (@SymbolicExprContext base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl).
- Local Notation SymbolicExprContextOk := (@SymbolicExprContextOk base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl base_type_code_lb op_code_bl op_code_lb).
- Local Notation prepend_prefix := (@prepend_prefix base_type_code op).
-
- Local Notation denote_symbolic_expr := (@denote_symbolic_expr base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op denote_op interp_base_type interp_op).
-
- Local Instance base_type_code_dec : DecidableRel (@eq base_type_code)
- := dec_rel_of_bool_dec_rel base_type_code_beq base_type_code_bl base_type_code_lb.
- Local Instance op_code_dec : DecidableRel (@eq op_code)
- := dec_rel_of_bool_dec_rel op_code_beq op_code_bl op_code_lb.
-
- Lemma interpf_flatten_binding_list T t x y v s
- (H : List.In (existT _ t (x, y)) (flatten_binding_list (var2:=interp_base_type) (t:=T) (symbolicify_smart_var v s) v))
- : fst x = y.
- Proof.
- revert dependent s; induction T;
- repeat first [ progress simpl in *
- | reflexivity
- | tauto
- | progress destruct_head or
- | progress inversion_sigma
- | progress inversion_prod
- | progress subst
- | rewrite List.in_app_iff in *
- | progress intros
- | solve [ eauto ] ].
- Qed.
-
- (*Lemma interpf_symbolize_exprf_injective
- s
- : forall G0 G1 t e0 e1 e0' e1'
- (HG0 : forall t x y, In (existT _ t (x, y)) G0 -> fst x = y)
- (HG1 : forall t x y, In (existT _ t (x, y)) G1 -> fst x = y)
- (Hwf0 : wff G0 (t:=t) e0 e0')
- (Hwf1 : wff G1 (t:=t) e1 e1')
- (Hs0 : symbolize_exprf e0 = Some s)
- (Hs1 : symbolize_exprf e1 = Some s),
- interpf interp_op e0' = interpf interp_op e1'.
- Proof.
- induction s; intros;
- destruct Hwf0;
- try (invert_one_expr e1; break_innermost_match; try exact I; intros);
- try (invert_one_expr e1'; break_innermost_match; try exact I; intros);
- try solve [ repeat first [ reflexivity
- | progress subst
- | progress destruct_head'_sig
- | progress destruct_head'_and
- | progress destruct_head'_prod
- | progress inversion_option
- | congruence
- | progress simpl in *
- | progress unfold option_map in *
- | progress inversion_wf_constr
- | break_innermost_match_hyps_step
- | match goal with
- | [ HG : forall t x y, In _ ?G -> fst x = y, H : In _ ?G |- _ ]
- => pose proof (HG _ _ _ H); progress subst
- end ] ].
- Focus 3.
- { simpl in *.
- Focus 3.
- try reflexivity;
- simpl in *.
- inversion_expr.
- .
- inversion_wf.
- move s at top; reverse.
- ->
-
-*)
-
-
-
- Local Arguments lookupb : simpl never.
- Local Arguments extendb : simpl never.
- Lemma interpf_csef G t e1 e2
- (HG : forall t x y, In (existT _ t (x, y)) G -> fst x = y)
- (m : @SymbolicExprContext (interp_flat_type_gen _))
- (HGm : forall t s v,
- lookupb m s = Some v
- -> forall k,
- List.In k (flatten_binding_list (@symbolicify_smart_var interp_base_type t v s) v)
- -> List.In k G)
- (Hm : forall t sv v,
- lookupb m sv = Some v
- -> denote_symbolic_expr m t sv = Some v)
- (Hwf : wff G e1 e2)
- : interpf interp_op (@csef interp_base_type t e1 m) = interpf interp_op e2.
- Proof.
- cbv beta in *.
- revert dependent m;
- induction Hwf; simpl; cbv [LetIn.Let_In symbolize_var]; intros; eauto;
- rewrite_hyp ?* by eauto; try reflexivity; eauto.
- (*{ break_match; break_match_hyps; try congruence; inversion_prod; inversion_option; subst;
- simpl; unfold LetIn.Let_In;
- [ match goal with
- | [ Hm : forall t e1 e2 s v, symbolize_exprf _ = _ -> _, H' : symbolize_exprf _ = _ |- _ ]
- => erewrite (Hm _ _ _ _ _ H') by eassumption
- end
- | rewrite_hyp !* by eauto
- | rewrite_hyp !* by eauto ];
- match goal with
- | [ H : context[interpf interp_op (csef _ _)] |- _ ] => erewrite H; [ reflexivity | | eauto | eauto ]
- end;
- intros *;
- try solve [ repeat first [ rewrite !(fun var => @lookupb_extendb_full flat_type _ symbolic_expr _ var _ SymbolicExprContextOk)
- | setoid_rewrite in_app_iff
- | progress break_innermost_match
- | progress subst
- | progress simpl in *
- | progress inversion_prod
- | progress inversion_option
- | progress destruct_head or
- | solve [ eauto using interpf_flatten_binding_list ]
- | progress intros ] ].
- admit.
- admit.
- admit. }*)
- Admitted.
-
- Lemma interpf_prepend_prefix t e prefix
- : interpf interp_op (@prepend_prefix _ t e prefix) = interpf interp_op e.
- Proof.
- induction prefix; simpl; [ reflexivity | unfold LetIn.Let_In; assumption ].
- Qed.
-
- Lemma interp_cse prefix t e1 e2
- (Hwf : wf e1 e2)
- : forall x, interp interp_op (@cse interp_base_type prefix t e1 empty) x = interp interp_op e2 x.
- Proof.
- destruct Hwf; simpl; intros.
- etransitivity; [ | eapply interpf_prepend_prefix ].
- eapply interpf_csef; eauto;
- [ ..
- | eapply wff_prepend_prefix; [ .. | solve [ eauto ] ] ].
- { intros; eapply interpf_flatten_binding_list; eassumption. }
- { admit. }
- { admit. }
- Admitted.
-
- Lemma InterpCSE t (e : Expr t)
- (prefix : forall var, list (sigT (fun t : flat_type => @exprf var t)))
- (*(Hlen : forall var1 var2, length (prefix var1) = length (prefix var2))
- (Hprefix : forall var1 var2 n t1 t2 e1 e2,
- nth_error (prefix var1) n = Some (existT _ t1 e1)
- -> nth_error (prefix var2) n = Some (existT _ t2 e2)
- -> exists pf : t1 = t2, wff nil (eq_rect _ exprf e1 _ pf) e2)*)
- (Hwf : Wf e)
- : forall x, Interp interp_op (@CSE t e prefix) x = Interp interp_op e x.
- Proof.
- apply interp_cse; auto.
- Qed.
-End symbolic.
-
-Hint Rewrite @InterpCSE using solve_wf_side_condition : reflective_interp.
diff --git a/src/Compilers/CommonSubexpressionEliminationProperties.v b/src/Compilers/CommonSubexpressionEliminationProperties.v
deleted file mode 100644
index b52afcb18..000000000
--- a/src/Compilers/CommonSubexpressionEliminationProperties.v
+++ /dev/null
@@ -1,189 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Coq.omega.Omega.
-Require Import Coq.Lists.List.
-Require Import Coq.FSets.FMapInterface.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Equality.
-Require Import Crypto.Compilers.CommonSubexpressionElimination.
-Require Import Crypto.Util.NatUtil.
-
-Local Open Scope list_scope.
-
-Local Open Scope ctype_scope.
-Section symbolic.
- (** Holds decidably-equal versions of raw expressions, for lookup. *)
- Context (base_type_code : Type)
- (op_code : Type)
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (op_code_beq : op_code -> op_code -> bool)
- (base_type_code_bl : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (op_code_bl : forall x y, op_code_beq x y = true -> x = y)
- (op_code_lb : forall x y, x = y -> op_code_beq x y = true)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (symbolize_op : forall s d, op s d -> op_code)
- (op_code_leb : op_code -> op_code -> bool)
- (base_type_leb : base_type_code -> base_type_code -> bool)
- (op_code_leb_total : forall x y, op_code_leb x y = true \/ op_code_leb y x = true)
- (base_type_leb_total : forall x y, base_type_leb x y = true \/ base_type_leb y x = true)
- (op_code_leb_antisymmetric : forall x y, op_code_leb x y = true -> op_code_leb y x = true -> op_code_beq x y = true)
- (base_type_leb_antisymmetric : forall x y, base_type_leb x y = true -> base_type_leb y x = true -> base_type_code_beq x y = true).
- Local Notation symbolic_expr := (symbolic_expr base_type_code op_code).
- Context (normalize_symbolic_op_arguments : op_code -> symbolic_expr -> symbolic_expr).
-
- Local Notation symbolic_expr_beq := (@symbolic_expr_beq base_type_code op_code base_type_code_beq op_code_beq).
- Local Notation symbolic_expr_lb := (@internal_symbolic_expr_dec_lb base_type_code op_code base_type_code_beq op_code_beq base_type_code_lb op_code_lb).
- Local Notation symbolic_expr_bl := (@internal_symbolic_expr_dec_bl base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op_code_bl).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Local Notation symbolic_expr_leb := (@symbolic_expr_leb base_type_code op_code base_type_code_beq op_code_beq op_code_leb base_type_leb).
- Local Notation flat_type_leb := (@flat_type_leb base_type_code base_type_code_beq base_type_leb).
- Local Notation flat_type_beq := (@flat_type_beq base_type_code base_type_code_beq).
- Local Notation flat_type_bl := (@internal_flat_type_dec_bl base_type_code base_type_code_beq base_type_code_bl).
- Local Notation flat_type_lb := (@internal_flat_type_dec_lb base_type_code base_type_code_beq base_type_code_lb).
-
- Lemma base_type_leb_reflexive x : base_type_leb x x = true.
- Proof using base_type_leb_total. destruct (base_type_leb_total x x); assumption. Qed.
-
- Lemma op_code_leb_reflexive x : op_code_leb x x = true.
- Proof using op_code_leb_total. destruct (op_code_leb_total x x); assumption. Qed.
-
- Theorem flat_type_leb_total : forall a1 a2, flat_type_leb a1 a2 = true \/ flat_type_leb a2 a1 = true.
- Proof using base_type_code_bl base_type_leb_total.
- induction a1, a2;
- repeat first [ progress simpl
- | progress subst
- | solve [ auto ]
- | match goal with
- | [ H : forall a2', ?leb ?a1 a2' = true \/ _ |- context[?leb ?a1 ?a2] ]
- => let H' := fresh in destruct (H a2) as [H'|H']; rewrite H'
- | [ H : flat_type_beq _ _ = true |- _ ] => apply flat_type_bl in H
- | [ |- context[flat_type_beq ?x ?y] ]
- => destruct (flat_type_beq x y) eqn:?
- end ].
- Qed.
-
- Theorem flat_type_leb_reflexive x : flat_type_leb x x = true.
- Proof using base_type_code_bl base_type_leb_total. destruct (flat_type_leb_total x x); assumption. Qed.
-
- Local Ltac rewrite_beq_leb_flat_type_op_code_reflexive :=
- repeat match goal with
- | [ H : flat_type_beq _ _ = true |- _ ] => apply flat_type_bl in H
- | [ H : op_code_beq _ _ = true |- _ ] => apply op_code_bl in H
- | [ H : symbolic_expr_beq _ _ = true |- _ ] => apply symbolic_expr_bl in H
- | [ H : context[flat_type_leb ?x ?x] |- _ ]
- => rewrite (flat_type_leb_reflexive x) in H
- | [ |- context[flat_type_leb ?x ?x] ]
- => rewrite (flat_type_leb_reflexive x)
- | [ H : context[flat_type_beq ?x ?x] |- _ ]
- => rewrite (flat_type_lb x x eq_refl) in H
- | [ |- context[flat_type_beq ?x ?x] ]
- => rewrite (flat_type_lb x x eq_refl)
- | [ H : context[op_code_leb ?x ?x] |- _ ]
- => rewrite (op_code_leb_reflexive x) in H
- | [ |- context[op_code_leb ?x ?x] ]
- => rewrite (op_code_leb_reflexive x)
- | [ H : context[op_code_beq ?x ?x] |- _ ]
- => rewrite (op_code_lb x x eq_refl) in H
- | [ |- context[op_code_beq ?x ?x] ]
- => rewrite (op_code_lb x x eq_refl)
- end.
-
- Theorem flat_type_leb_antisymmetric : forall a1 a2, flat_type_leb a1 a2 = true -> flat_type_leb a2 a1 = true -> flat_type_beq a1 a2 = true.
- Proof using base_type_code_bl base_type_code_lb base_type_leb_antisymmetric base_type_leb_total.
- induction a1, a2;
- repeat first [ progress simpl
- | progress subst
- | solve [ auto ]
- | progress rewrite ?andb_true_r, ?orb_false_r
- | progress rewrite_beq_leb_flat_type_op_code_reflexive
- | match goal with
- | [ |- context[flat_type_beq ?x ?y] ]
- => destruct (flat_type_beq x y) eqn:?
- | [ H : forall a2, ?leb ?x a2 = true -> ?leb a2 ?x = true -> _, H0 : ?leb ?x ?a2' = true, H1 : ?leb ?a2' ?x = true |- _ ]
- => specialize (H _ H0 H1)
- end
- | progress intros ].
- Qed.
-
- Theorem symbolic_expr_leb_total : forall a1 a2, symbolic_expr_leb a1 a2 = true \/ symbolic_expr_leb a2 a1 = true.
- Proof using base_type_code_bl base_type_code_lb base_type_leb_total op_code_bl op_code_lb op_code_leb_total.
- induction a1, a2;
- repeat first [ rewrite !PeanoNat.Nat.leb_le
- | progress subst
- | progress simpl
- | solve [ auto ]
- | omega
- | progress rewrite_beq_leb_flat_type_op_code_reflexive
- | match goal with
- | [ |- context[flat_type_beq ?x ?y] ]
- => destruct (flat_type_beq x y) eqn:?
- | [ |- context[op_code_beq ?x ?y] ]
- => destruct (op_code_beq x y) eqn:?
- | [ |- context[symbolic_expr_beq ?x ?y] ]
- => destruct (symbolic_expr_beq x y) eqn:?
- | [ H : forall a2', ?leb ?a1 a2' = true \/ _ |- context[?leb ?a1 ?a2] ]
- => let H' := fresh in destruct (H a2) as [H'|H']; rewrite H'
- | [ |- context[flat_type_leb ?a1 ?a2] ]
- => let H' := fresh in destruct (flat_type_leb_total a1 a2) as [H'|H']; rewrite H'
- | [ |- context[op_code_leb ?a1 ?a2] ]
- => let H' := fresh in destruct (op_code_leb_total a1 a2) as [H'|H']; rewrite H'
- end ].
- Qed.
-
- Theorem symbolic_expr_leb_reflexive x : symbolic_expr_leb x x = true.
- Proof using base_type_code_bl base_type_code_lb base_type_leb_total op_code_bl op_code_lb op_code_leb_total. destruct (symbolic_expr_leb_total x x); assumption. Qed.
-
- Local Ltac rewrite_beq_leb_symbolic_expr_reflexive :=
- repeat match goal with
- | [ H : symbolic_expr_beq _ _ = true |- _ ] => apply symbolic_expr_bl in H
- | [ H : context[symbolic_expr_leb ?x ?x] |- _ ]
- => rewrite (symbolic_expr_leb_reflexive x) in H
- | [ |- context[symbolic_expr_leb ?x ?x] ]
- => rewrite (symbolic_expr_leb_reflexive x)
- | [ H : context[symbolic_expr_beq ?x ?x] |- _ ]
- => rewrite (symbolic_expr_lb x x eq_refl) in H
- | [ |- context[symbolic_expr_beq ?x ?x] ]
- => rewrite (symbolic_expr_lb x x eq_refl)
- end.
-
- Theorem symbolic_expr_leb_antisymmetric : forall a1 a2, symbolic_expr_leb a1 a2 = true -> symbolic_expr_leb a2 a1 = true -> symbolic_expr_beq a1 a2 = true.
- Proof using base_type_code_bl base_type_code_lb base_type_leb_antisymmetric base_type_leb_total op_code_bl op_code_lb op_code_leb_antisymmetric op_code_leb_total.
- induction a1, a2;
- repeat first [ rewrite !PeanoNat.Nat.leb_le
- | progress subst
- | progress simpl
- | solve [ auto ]
- | omega
- | progress rewrite ?andb_true_r, ?orb_false_r
- | progress rewrite_beq_leb_flat_type_op_code_reflexive
- | progress rewrite_beq_leb_symbolic_expr_reflexive
- | match goal with
- | [ |- context[flat_type_beq ?x ?y] ]
- => destruct (flat_type_beq x y) eqn:?
- | [ |- context[op_code_beq ?x ?y] ]
- => destruct (op_code_beq x y) eqn:?
- | [ |- context[symbolic_expr_beq ?x ?y] ]
- => destruct (symbolic_expr_beq x y) eqn:?
- | [ |- context[nat_beq ?x ?x] ]
- => rewrite (internal_nat_dec_lb x x eq_refl)
- | [ |- context[flat_type_leb ?a1 ?a2] ]
- => let H' := fresh in destruct (flat_type_leb_total a1 a2) as [H'|H']; rewrite H'
- | [ |- context[op_code_leb ?a1 ?a2] ]
- => let H' := fresh in destruct (op_code_leb_total a1 a2) as [H'|H']; rewrite H'
- | [ H : ?x <= ?y, H' : ?y <= ?x |- _ ] => assert (x = y) by omega; clear H H'
- | [ H : forall a2, ?leb ?x a2 = true -> ?leb a2 ?x = true -> _, H0 : ?leb ?x ?a2' = true, H1 : ?leb ?a2' ?x = true |- _ ]
- => specialize (H _ H0 H1)
- | [ H0 : flat_type_leb ?x' ?y' = true, H1 : flat_type_leb ?y' ?x' = true |- _ ]
- => pose proof (flat_type_leb_antisymmetric _ _ H0 H1); clear H0 H1
- | [ H : forall x y, ?leb x y = true -> ?leb y x = true -> _, H0 : ?leb ?x' ?y' = true, H1 : ?leb ?y' ?x' = true |- _ ]
- => pose proof (H _ _ H0 H1); clear H0 H1
- end
- | progress intros ].
- Qed.
-End symbolic.
diff --git a/src/Compilers/CommonSubexpressionEliminationWf.v b/src/Compilers/CommonSubexpressionEliminationWf.v
deleted file mode 100644
index 8d7bf2727..000000000
--- a/src/Compilers/CommonSubexpressionEliminationWf.v
+++ /dev/null
@@ -1,223 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.AListContext.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Equality.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.CommonSubexpressionElimination.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.UniquePose.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Decidable.
-
-Section symbolic.
- Context (base_type_code : Type)
- (op_code : Type)
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (op_code_beq : op_code -> op_code -> bool)
- (base_type_code_bl : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (op_code_bl : forall x y, op_code_beq x y = true -> x = y)
- (op_code_lb : forall x y, x = y -> op_code_beq x y = true)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (symbolize_op : forall s d, op s d -> op_code).
- Local Notation symbolic_expr := (symbolic_expr base_type_code op_code).
- Context (normalize_symbolic_op_arguments : op_code -> symbolic_expr -> symbolic_expr)
- (inline_symbolic_expr_in_lookup : bool).
-
- Local Notation symbolic_expr_beq := (@symbolic_expr_beq base_type_code op_code base_type_code_beq op_code_beq).
- Local Notation symbolic_expr_lb := (@internal_symbolic_expr_dec_lb base_type_code op_code base_type_code_beq op_code_beq base_type_code_lb op_code_lb).
- Local Notation symbolic_expr_bl := (@internal_symbolic_expr_dec_bl base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op_code_bl).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Local Notation symbolicify_smart_var := (@symbolicify_smart_var base_type_code op_code).
- Local Notation symbolize_exprf := (@symbolize_exprf base_type_code op_code op symbolize_op).
- Local Notation norm_symbolize_exprf := (@norm_symbolize_exprf base_type_code op_code op symbolize_op normalize_symbolic_op_arguments).
- Local Notation csef := (@csef base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup).
- Local Notation cse := (@cse base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup).
- Local Notation CSE := (@CSE base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl op symbolize_op normalize_symbolic_op_arguments inline_symbolic_expr_in_lookup).
- Local Notation SymbolicExprContext := (@SymbolicExprContext base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl).
- Local Notation SymbolicExprContextOk := (@SymbolicExprContextOk base_type_code op_code base_type_code_beq op_code_beq base_type_code_bl base_type_code_lb op_code_bl op_code_lb).
- Local Notation prepend_prefix := (@prepend_prefix base_type_code op).
-
- Local Instance base_type_code_dec : DecidableRel (@eq base_type_code)
- := dec_rel_of_bool_dec_rel base_type_code_beq base_type_code_bl base_type_code_lb.
- Local Instance op_code_dec : DecidableRel (@eq op_code)
- := dec_rel_of_bool_dec_rel op_code_beq op_code_bl op_code_lb.
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}.
-
- Lemma wff_symbolize_exprf G t e1 e2
- (HG : forall t x y, List.In (existT _ t (x, y)) G -> snd x = snd y)
- (Hwf : wff G e1 e2)
- : @symbolize_exprf var1 t e1 = @symbolize_exprf var2 t e2.
- Proof.
- induction Hwf; simpl; erewrite_hyp ?* by eassumption; reflexivity.
- Qed.
-
- Lemma wff_norm_symbolize_exprf G t e1 e2
- (HG : forall t x y, List.In (existT _ t (x, y)) G -> snd x = snd y)
- (Hwf : wff G e1 e2)
- : @norm_symbolize_exprf var1 t e1 = @norm_symbolize_exprf var2 t e2.
- Proof.
- unfold norm_symbolize_exprf; erewrite wff_symbolize_exprf by eassumption; reflexivity.
- Qed.
-
- Local Arguments lookupb : simpl never.
- Local Arguments extendb : simpl never.
- Lemma wff_csef G G' t e1 e2
- (m1 : @SymbolicExprContext (interp_flat_type var1))
- (m2 : @SymbolicExprContext (interp_flat_type var2))
- (Hlen : length m1 = length m2)
- (Hm1m2None : forall t v, lookupb t m1 v = None <-> lookupb t m2 v = None)
- (Hm1m2Some : forall t v sv1 sv2,
- lookupb t m1 v = Some sv1
- -> lookupb t m2 v = Some sv2
- -> forall k,
- List.In k (flatten_binding_list
- (t:=t)
- (symbolicify_smart_var sv1 v)
- (symbolicify_smart_var sv2 v))
- -> List.In k G)
- (HG : forall t x y, List.In (existT _ t (x, y)) G -> snd x = snd y)
- (HGG' : forall t x x', List.In (existT _ t (x, x')) G -> List.In (existT _ t (fst x, fst x')) G')
- (Hwf : wff G e1 e2)
- : wff G' (@csef var1 t e1 m1) (@csef var2 t e2 m2).
- Proof.
- revert dependent m1; revert m2; revert dependent G'.
- induction Hwf; simpl; intros G' HGG' m2 m1 Hlen Hm1m2None Hm1m2Some; try constructor; auto.
- { erewrite wff_norm_symbolize_exprf by eassumption.
- break_innermost_match;
- try match goal with
- | [ H : lookupb ?m1 ?x = Some ?k, H' : lookupb ?m2 ?x = None |- _ ]
- => apply Hm1m2None in H'; congruence
- end;
- lazymatch goal with
- | [ |- wff _ (LetIn _ _) (LetIn _ _) ]
- => constructor; intros; auto; []
- | _ => idtac
- end;
- match goal with H : _ |- _ => apply H end;
- try solve [ repeat first [ progress unfold symbolize_var
- | rewrite Hlen
- | progress subst
- | setoid_rewrite length_extendb
- | setoid_rewrite List.in_app_iff
- | progress destruct_head' or
- | solve [ eauto ]
- | progress intros
- | match goal with
- | [ H : List.In _ (flatten_binding_list (symbolicify_smart_var ?x1 ?v) (symbolicify_smart_var ?x2 ?v)) |- _ ]
- => solve [ destruct (flatten_binding_list_SmartVarfMap2_pair_In_split H); eauto ]
- | [ H : List.In _ (flatten_binding_list (symbolicify_smart_var ?x1 ?v) (symbolicify_smart_var ?x2 ?v)) |- _ ]
- => exact (flatten_binding_list_SmartVarfMap2_pair_same_in_eq2 H)
- | [ H : context[lookupb (extendb _ _ _) _] |- _ ]
- => rewrite (fun var => @lookupb_extendb_full flat_type _ symbolic_expr _ var _ SymbolicExprContextOk) in H
- end
- | rewrite !(fun var => @lookupb_extendb_full flat_type _ symbolic_expr _ var _ SymbolicExprContextOk)
- | break_innermost_match_step
- | break_innermost_match_hyps_step
- | progress simpl in *
- | solve [ intuition (eauto || congruence) ]
- | match goal with
- | [ H : forall t x y, _ |- _ ] => specialize (fun t x0 x1 y0 y1 => H t (x0, x1) (y0, y1)); cbn [fst snd] in H
- | [ H : In (existT _ ?t (?x, ?x')) (flatten_binding_list (symbolicify_smart_var _ _) (symbolicify_smart_var _ _)),
- Hm1m2Some : forall t v sv1 sv2, _ -> _ -> forall k', In k' (flatten_binding_list _ _) -> In k' ?G |- _ ]
- => is_var x; is_var x';
- lazymatch goal with
- | [ H : In (existT _ t ((fst x, _), (fst x', _))) G |- _ ] => fail
- | _ => let H' := fresh in
- refine (let H' := flatten_binding_list_SmartVarfMap2_pair_in_generalize2 H _ _ in _);
- destruct H' as [? [? H']];
- eapply Hm1m2Some in H'; [ | eassumption.. ]
- end
- end ] ].
- repeat first [ progress unfold symbolize_var
- | rewrite Hlen
- | progress subst
- | setoid_rewrite length_extendb
- | setoid_rewrite List.in_app_iff
- | progress destruct_head' or
- | solve [ eauto ]
- | progress intros ].
- (** FIXME: This actually isn't true, because the symbolic
- expr stored in G might not be the same as the one in the
- expression tree, when the one in the expression tree is a
- fresh var *)
- admit. }
- Admitted.
-
- Lemma wff_prepend_prefix {var1' var2'} prefix1 prefix2 G t e1 e2
- (Hlen : length prefix1 = length prefix2)
- (Hprefix : forall n t1 t2 e1 e2,
- nth_error prefix1 n = Some (existT _ t1 e1)
- -> nth_error prefix2 n = Some (existT _ t2 e2)
- -> exists pf : t1 = t2, wff nil (eq_rect _ exprf e1 _ pf) e2)
- (Hwf : wff G e1 e2)
- : wff G (@prepend_prefix var1' t e1 prefix1) (@prepend_prefix var2' t e2 prefix2).
- Proof.
- revert dependent G; revert dependent prefix2; induction prefix1, prefix2; simpl; intros Hlen Hprefix G Hwf; try congruence.
- { pose proof (Hprefix 0) as H0; specialize (fun n => Hprefix (S n)).
- destruct_head sigT; simpl in *.
- specialize (H0 _ _ _ _ eq_refl eq_refl); destruct_head ex; subst; simpl in *.
- constructor.
- { eapply wff_in_impl_Proper; [ eassumption | simpl; tauto ]. }
- { intros.
- apply IHprefix1; try congruence; auto.
- eapply wff_in_impl_Proper; [ eassumption | simpl; intros; rewrite in_app_iff; auto ]. } }
- Qed.
-
- Lemma wf_cse prefix1 prefix2 t e1 e2 (Hwf : wf e1 e2)
- (Hlen : length prefix1 = length prefix2)
- (Hprefix : forall n t1 t2 e1 e2,
- nth_error prefix1 n = Some (existT _ t1 e1)
- -> nth_error prefix2 n = Some (existT _ t2 e2)
- -> exists pf : t1 = t2, wff nil (eq_rect _ exprf e1 _ pf) e2)
- : wf (@cse var1 prefix1 t e1 empty) (@cse var2 prefix2 t e2 empty).
- Proof.
- destruct Hwf; simpl; constructor; intros.
- lazymatch goal with
- | [ |- wff (flatten_binding_list (t:=?src) ?x ?y) (csef _ (extendb _ ?n ?v)) (csef _ (extendb _ ?n' ?v')) ]
- => unify n n';
- apply wff_csef with (G:=flatten_binding_list (t:=src) (symbolicify_smart_var v n) (symbolicify_smart_var v' n'))
- end.
- { setoid_rewrite length_extendb; reflexivity. }
- { intros; rewrite !(fun var => @lookupb_extendb_full flat_type _ symbolic_expr _ var _ SymbolicExprContextOk).
- break_innermost_match; subst; simpl; intuition (eauto || congruence). }
- { intros *; rewrite !(fun var => @lookupb_extendb_full flat_type _ symbolic_expr _ var _ SymbolicExprContextOk).
- break_innermost_match; subst; simpl; try setoid_rewrite lookupb_empty; eauto using SymbolicExprContextOk; try congruence. }
- { intros *; intro H'; exact (flatten_binding_list_SmartVarfMap2_pair_same_in_eq2 H'). }
- { intros *; intro H'; destruct (flatten_binding_list_SmartVarfMap2_pair_In_split H'); eauto. }
- { apply wff_prepend_prefix; auto. }
- Qed.
- End with_var.
-
- Lemma Wf_CSE t (e : Expr t)
- (prefix : forall var, list (sigT (fun t : flat_type => @exprf var t)))
- (Hlen : forall var1 var2, length (prefix var1) = length (prefix var2))
- (Hprefix : forall var1 var2 n t1 t2 e1 e2,
- nth_error (prefix var1) n = Some (existT _ t1 e1)
- -> nth_error (prefix var2) n = Some (existT _ t2 e2)
- -> exists pf : t1 = t2, wff nil (eq_rect _ exprf e1 _ pf) e2)
- (Hwf : Wf e)
- : Wf (@CSE t e prefix).
- Proof.
- intros var1 var2; apply wf_cse; eauto.
- Qed.
-End symbolic.
-
-Hint Resolve Wf_CSE : wf.
diff --git a/src/Compilers/Conversion.v b/src/Compilers/Conversion.v
deleted file mode 100644
index 29874c96f..000000000
--- a/src/Compilers/Conversion.v
+++ /dev/null
@@ -1,110 +0,0 @@
-(** * Convert between interpretations of types *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Map.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-
-Local Open Scope expr_scope.
-
-Section language.
- Context (base_type_code : Type).
- Context (op : flat_type base_type_code -> flat_type base_type_code -> Type).
- Section map.
- Context {var1 var2 : base_type_code -> Type}.
- Context (f_var12 : forall t, var1 t -> var2 t)
- (f_var21 : forall t, var2 t -> var1 t).
-
- Fixpoint mapf
- {t}
- (e : @exprf base_type_code op var1 t)
- : @exprf base_type_code op var2 t
- := match e in @exprf _ _ _ t return @exprf _ _ _ t with
- | TT => TT
- | Var _ x => Var (f_var12 _ x)
- | Op _ _ op args => Op op (@mapf _ args)
- | LetIn _ ex _ eC => LetIn (@mapf _ ex)
- (fun x => @mapf _ (eC (mapf_interp_flat_type f_var21 x)))
- | Pair _ ex _ ey => Pair (@mapf _ ex)
- (@mapf _ ey)
- end.
-
- Definition map {t} (e : @expr base_type_code op var1 t)
- : @expr base_type_code op var2 t
- := match e with
- | Abs _ _ f => Abs (fun x => mapf (f (mapf_interp_flat_type f_var21 x)))
- end.
- End map.
-
- Section mapf_id.
- Context (functional_extensionality : forall {A B} (f g : A -> B), (forall x, f x = g x) -> f = g)
- {var : base_type_code -> Type}.
-
- Lemma mapf_idmap_ext {t} e
- : @mapf var var
- (fun _ x => x) (fun _ x => x)
- t e
- = e.
- Proof using functional_extensionality.
- induction e;
- repeat match goal with
- | _ => reflexivity
- | _ => progress simpl in *
- | _ => rewrite_hyp !*
- | _ => apply (f_equal2 (fun x f => LetIn x f))
- | _ => solve [ eauto ]
- | _ => apply functional_extensionality; intro
- end.
- clear e IHe H.
- revert dependent tC; induction tx; simpl; [ reflexivity | reflexivity | ]; intros.
- destruct x as [x0 x1]; simpl in *.
- lazymatch goal with
- | [ |- ?e0 (?x0', ?x1')%core = _ ]
- => rewrite (IHtx1 x0 _ (fun x0'' => e0 (x0'', x1')%core)); cbv beta in *
- end.
- lazymatch goal with
- | [ |- ?e0 (?x0', ?x1')%core = _ ]
- => rewrite (IHtx2 x1 _ (fun x1'' => e0 (x0', x1'')%core)); cbv beta in *
- end.
- reflexivity.
- Qed.
- End mapf_id.
-
- Section mapf_id_interp.
- Context {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst)
- (f_var12 f_var21 : forall t, interp_base_type t -> interp_base_type t)
- (f_var12_id : forall t x, f_var12 t x = x)
- (f_var21_id : forall t x, f_var21 t x = x).
-
- Lemma mapf_idmap {t} e
- : interpf interp_op
- (@mapf _ _
- f_var12 f_var21
- t e)
- = interpf interp_op e.
- Proof using f_var12_id f_var21_id.
- induction e;
- repeat match goal with
- | _ => progress unfold LetIn.Let_In
- | _ => reflexivity
- | _ => progress simpl in *
- | _ => rewrite_hyp !*
- | _ => apply (f_equal2 (fun x f => LetIn x f))
- | _ => solve [ eauto ]
- end.
- clear H IHe.
- generalize (interpf interp_op e); intro x; clear e.
- revert dependent tC; induction tx; simpl;
- [ intros; rewrite_hyp ?*; reflexivity | reflexivity | ]; intros.
- destruct x as [x0 x1]; simpl in *.
- lazymatch goal with
- | [ |- interpf _ (?e0 (?x0', ?x1')%core) = _ ]
- => rewrite (IHtx1 x0 _ (fun x0'' => e0 (x0'', x1')%core)); cbv beta in *
- end.
- lazymatch goal with
- | [ |- interpf _ (?e0 (?x0', ?x1')%core) = _ ]
- => apply (IHtx2 x1 _ (fun x1'' => e0 (x0', x1'')%core)); cbv beta in *
- end.
- Qed.
- End mapf_id_interp.
-End language.
diff --git a/src/Compilers/CountLets.v b/src/Compilers/CountLets.v
deleted file mode 100644
index 4810162c8..000000000
--- a/src/Compilers/CountLets.v
+++ /dev/null
@@ -1,66 +0,0 @@
-(** * Counts how many binders there are *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation Expr := (@Expr base_type_code op).
-
- Fixpoint count_pairs (t : flat_type) : nat
- := match t with
- | Tbase _ => 1
- | Unit => 0
- | Prod A B => count_pairs A + count_pairs B
- end%nat.
-
- Section with_var.
- Context {var : base_type_code -> Type}
- (mkVar : forall t, var t).
-
- Local Notation exprf := (@exprf base_type_code op var).
- Local Notation expr := (@expr base_type_code op var).
-
- Section gen.
- Context (count_type_let : flat_type -> nat).
- Context (count_type_abs : flat_type -> nat).
-
- Fixpoint count_lets_genf {t} (e : exprf t) : nat
- := match e with
- | LetIn tx _ _ eC
- => count_type_let tx + @count_lets_genf _ (eC (SmartValf var mkVar tx))
- | Op _ _ _ e => @count_lets_genf _ e
- | Pair _ ex _ ey => @count_lets_genf _ ex + @count_lets_genf _ ey
- | _ => 0
- end.
- Definition count_lets_gen {t} (e : expr t) : nat
- := match e with
- | Abs tx _ f => count_type_abs tx + @count_lets_genf _ (f (SmartValf _ mkVar tx))
- end.
- End gen.
-
- Definition count_let_bindersf {t} (e : exprf t) : nat
- := count_lets_genf count_pairs e.
- Definition count_letsf {t} (e : exprf t) : nat
- := count_lets_genf (fun _ => 1) e.
- Definition count_let_binders {t} (e : expr t) : nat
- := count_lets_gen count_pairs (fun _ => 0) e.
- Definition count_lets {t} (e : expr t) : nat
- := count_lets_gen (fun _ => 1) (fun _ => 0) e.
- Definition count_binders {t} (e : expr t) : nat
- := count_lets_gen count_pairs count_pairs e.
- End with_var.
-
- Definition CountLetsGen (count_type_let : flat_type -> nat) (count_type_abs : flat_type -> nat) {t} (e : Expr t) : nat
- := count_lets_gen (fun _ => tt) count_type_let count_type_abs (e _).
- Definition CountLetBinders {t} (e : Expr t) : nat
- := count_let_binders (fun _ => tt) (e _).
- Definition CountLets {t} (e : Expr t) : nat
- := count_lets (fun _ => tt) (e _).
- Definition CountBinders {t} (e : Expr t) : nat
- := count_binders (fun _ => tt) (e _).
-End language.
diff --git a/src/Compilers/Equality.v b/src/Compilers/Equality.v
deleted file mode 100644
index 2b54a7892..000000000
--- a/src/Compilers/Equality.v
+++ /dev/null
@@ -1,90 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.FixCoqMistakes.
-
-Section language.
- Context (base_type_code : Type)
- (eq_base_type_code : base_type_code -> base_type_code -> bool)
- (base_type_code_bl : forall x y, eq_base_type_code x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> eq_base_type_code x y = true).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
-
- Fixpoint flat_type_beq (X Y : flat_type) {struct X} : bool
- := match X, Y with
- | Tbase T, Tbase T0 => eq_base_type_code T T0
- | Unit, Unit => true
- | Prod A B, Prod A0 B0 => (flat_type_beq A A0 && flat_type_beq B B0)%bool
- | Tbase _, _
- | Prod _ _, _
- | Unit, _
- => false
- end.
- Local Ltac t :=
- repeat match goal with
- | _ => intro
- | _ => reflexivity
- | _ => assumption
- | _ => progress simpl in *
- | _ => solve [ eauto with nocore ]
- | [ H : False |- _ ] => exfalso; assumption
- | [ H : false = true |- _ ] => apply Bool.diff_false_true in H
- | [ |- Prod _ _ = Prod _ _ ] => apply f_equal2
- | [ |- Arrow _ _ = Arrow _ _ ] => apply f_equal2
- | [ |- Tbase _ = Tbase _ ] => apply f_equal
- | [ H : forall Y, _ = true -> _ = Y |- _ = ?Y' ]
- => is_var Y'; apply H; solve [ t ]
- | [ H : forall X Y, X = Y -> _ = true |- _ = true ]
- => eapply H; solve [ t ]
- | [ H : true = true |- _ ] => clear H
- | [ H : andb ?x ?y = true |- _ ]
- => destruct x, y; simpl in H; solve [ t ]
- | [ H : andb ?x ?y = true |- _ ]
- => destruct x eqn:?; simpl in H
- | [ H : ?f ?x = true |- _ ] => destruct (f x); solve [ t ]
- | [ H : ?x = true |- andb _ ?x = true ]
- => destruct x
- | [ |- andb ?x _ = true ]
- => cut (x = true); [ destruct x; simpl | ]
- end.
- Lemma flat_type_dec_bl X : forall Y, flat_type_beq X Y = true -> X = Y.
- Proof. clear base_type_code_lb; induction X, Y; t. Defined.
- Lemma flat_type_dec_lb X : forall Y, X = Y -> flat_type_beq X Y = true.
- Proof. clear base_type_code_bl; intros Y **; subst Y; induction X; t. Defined.
- Definition flat_type_eq_dec (X Y : flat_type) : {X = Y} + {X <> Y}
- := match Sumbool.sumbool_of_bool (flat_type_beq X Y) with
- | left pf => left (flat_type_dec_bl _ _ pf)
- | right pf => right (fun pf' => let pf'' := eq_sym (flat_type_dec_lb _ _ pf') in
- Bool.diff_true_false (eq_trans pf'' pf))
- end.
- Definition type_beq (X Y : type) : bool
- := match X, Y with
- | Arrow A B, Arrow A0 B0 => (flat_type_beq A A0 && flat_type_beq B B0)%bool
- end.
- Lemma type_dec_bl X : forall Y, type_beq X Y = true -> X = Y.
- Proof. clear base_type_code_lb; pose proof flat_type_dec_bl; induction X, Y; t. Defined.
- Lemma type_dec_lb X : forall Y, X = Y -> type_beq X Y = true.
- Proof. clear base_type_code_bl; pose proof flat_type_dec_lb; intros Y **; subst Y; induction X; t. Defined.
- Definition type_eq_dec (X Y : type) : {X = Y} + {X <> Y}
- := match Sumbool.sumbool_of_bool (type_beq X Y) with
- | left pf => left (type_dec_bl _ _ pf)
- | right pf => right (fun pf' => let pf'' := eq_sym (type_dec_lb _ _ pf') in
- Bool.diff_true_false (eq_trans pf'' pf))
- end.
-End language.
-
-Lemma dec_eq_flat_type {base_type_code} `{DecidableRel (@eq base_type_code)}
- : DecidableRel (@eq (flat_type base_type_code)).
-Proof.
- repeat intro; hnf; decide equality; apply dec; auto.
-Defined.
-Hint Extern 1 (Decidable (@eq (flat_type ?base_type_code) ?x ?y))
-=> simple apply (@dec_eq_flat_type base_type_code) : typeclass_instances.
-Lemma dec_eq_type {base_type_code} `{DecidableRel (@eq base_type_code)}
- : DecidableRel (@eq (type base_type_code)).
-Proof.
- repeat intro; hnf; decide equality; apply dec; typeclasses eauto.
-Defined.
-Hint Extern 1 (Decidable (@eq (type ?base_type_code) ?x ?y))
-=> simple apply (@dec_eq_type base_type_code) : typeclass_instances.
diff --git a/src/Compilers/Eta.v b/src/Compilers/Eta.v
deleted file mode 100644
index 9ca778f15..000000000
--- a/src/Compilers/Eta.v
+++ /dev/null
@@ -1,75 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Local Notation Expr := (@Expr base_type_code op).
- Section with_var.
- Context {var : base_type_code -> Type}.
- Local Notation exprf := (@exprf base_type_code op var).
- Local Notation expr := (@expr base_type_code op var).
-
- Section gen_flat_type.
- Context (eta : forall {A B}, A * B -> A * B).
- Fixpoint interp_flat_type_eta_gen {t T} : (interp_flat_type var t -> T) -> interp_flat_type var t -> T
- := match t return (interp_flat_type var t -> T) -> interp_flat_type var t -> T with
- | Tbase T => fun f => f
- | Unit => fun f => f
- | Prod A B
- => fun f x
- => let '(a, b) := eta _ _ x in
- @interp_flat_type_eta_gen
- A _
- (fun a' => @interp_flat_type_eta_gen B _ (fun b' => f (a', b')) b)
- a
- end.
-
- Section gen_type.
- Context (exprf_eta : forall {t} (e : exprf t), exprf t).
- Definition expr_eta_gen {t} (e : expr t) : expr (Arrow (domain t) (codomain t))
- := Abs (interp_flat_type_eta_gen (fun x => exprf_eta _ (invert_Abs e x))).
- End gen_type.
-
- Fixpoint exprf_eta_gen {t} (e : exprf t) : exprf t
- := match e in Syntax.exprf _ _ t return exprf t with
- | TT => TT
- | Var t v => Var v
- | Op t1 tR opc args => Op opc (@exprf_eta_gen _ args)
- | LetIn tx ex tC eC
- => LetIn (@exprf_eta_gen _ ex)
- (interp_flat_type_eta_gen (fun x => @exprf_eta_gen _ (eC x)))
- | Pair tx ex ty ey => Pair (@exprf_eta_gen _ ex) (@exprf_eta_gen _ ey)
- end.
- End gen_flat_type.
-
- Definition interp_flat_type_eta {t T}
- := @interp_flat_type_eta_gen (fun _ _ x => x) t T.
- Definition interp_flat_type_eta' {t T}
- := @interp_flat_type_eta_gen (fun _ _ x => (fst x, snd x)) t T.
- Definition exprf_eta {t}
- := @exprf_eta_gen (fun _ _ x => x) t.
- Definition exprf_eta' {t}
- := @exprf_eta_gen (fun _ _ x => (fst x, snd x)) t.
- Definition expr_eta {t}
- := @expr_eta_gen (fun _ _ x => x) (@exprf_eta) t.
- Definition expr_eta' {t}
- := @expr_eta_gen (fun _ _ x => (fst x, snd x)) (@exprf_eta') t.
- End with_var.
- Definition ExprEtaGen eta {t} (e : Expr t) : Expr (Arrow (domain t) (codomain t))
- := fun var => expr_eta_gen eta (@exprf_eta_gen var eta) (e var).
- Definition ExprEta {t} (e : Expr t) : Expr (Arrow (domain t) (codomain t))
- := fun var => expr_eta (e var).
- Definition ExprEta' {t} (e : Expr t) : Expr (Arrow (domain t) (codomain t))
- := fun var => expr_eta' (e var).
-End language.
-(* put these outside the section so the argument order lines up with
- [interp] and [Interp] *)
-Definition interp_eta {base_type_code interp_base_type op} interp_op
- {t} (e : @expr base_type_code op interp_base_type t)
- : interp_type interp_base_type t
- := interp_flat_type_eta (interp interp_op e).
-Definition InterpEta {base_type_code interp_base_type op} interp_op
- {t} (e : @Expr base_type_code op t)
- : interp_type interp_base_type t
- := interp_eta interp_op (e _).
diff --git a/src/Compilers/EtaInterp.v b/src/Compilers/EtaInterp.v
deleted file mode 100644
index 2b6bd9b86..000000000
--- a/src/Compilers/EtaInterp.v
+++ /dev/null
@@ -1,118 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Eta.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Section language.
- Context {base_type_code : Type}
- {interp_base_type : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}.
-
- Local Notation exprf := (@exprf base_type_code op interp_base_type).
-
- Local Ltac t_step :=
- match goal with
- | _ => reflexivity
- | _ => progress simpl in *
- | _ => intro
- | _ => progress break_match
- | _ => progress destruct_head prod
- | _ => progress cbv [LetIn.Let_In]
- | [ H : _ |- _ ] => rewrite H
- | _ => progress autorewrite with core
- | [ H : forall A B x, ?f A B x = x, H' : context[?f _ _ _] |- _ ]
- => rewrite H in H'
- | _ => progress unfold interp_flat_type_eta, interp_flat_type_eta', exprf_eta, exprf_eta', expr_eta, expr_eta'
- end.
- Local Ltac t := repeat t_step.
-
- Section gen_flat_type.
- Context (eta : forall {A B}, A * B -> A * B)
- (eq_eta : forall A B x, @eta A B x = x).
- Lemma eq_interp_flat_type_eta_gen {var t T f} x
- : @interp_flat_type_eta_gen base_type_code var eta t T f x = f x.
- Proof using eq_eta. induction t; t. Qed.
-
- (* Local *) Hint Rewrite @eq_interp_flat_type_eta_gen.
-
- Section gen_type.
- Context (exprf_eta : forall {t} (e : exprf t), exprf t)
- (eq_interp_exprf_eta : forall t e, interpf (@interp_op) (@exprf_eta t e) = interpf (@interp_op) e).
- Lemma interp_expr_eta_gen {t e}
- : forall x,
- interp (@interp_op) (expr_eta_gen eta exprf_eta (t:=t) e) x = interp (@interp_op) e x.
- Proof using Type*. t. Qed.
- End gen_type.
- (* Local *) Hint Rewrite @interp_expr_eta_gen.
-
- Lemma interpf_exprf_eta_gen {t e}
- : interpf (@interp_op) (exprf_eta_gen eta (t:=t) e) = interpf (@interp_op) e.
- Proof using eq_eta. induction e; t. Qed.
-
- Lemma InterpExprEtaGen {t e}
- : forall x, Interp (@interp_op) (ExprEtaGen eta (t:=t) e) x = Interp (@interp_op) e x.
- Proof using eq_eta. apply interp_expr_eta_gen; intros; apply interpf_exprf_eta_gen. Qed.
- End gen_flat_type.
- (* Local *) Hint Rewrite @eq_interp_flat_type_eta_gen.
- (* Local *) Hint Rewrite @interp_expr_eta_gen.
- (* Local *) Hint Rewrite @interpf_exprf_eta_gen.
-
- Lemma eq_interp_flat_type_eta {var t T f} x
- : @interp_flat_type_eta base_type_code var t T f x = f x.
- Proof using Type. t. Qed.
- (* Local *) Hint Rewrite @eq_interp_flat_type_eta.
- Lemma eq_interp_flat_type_eta' {var t T f} x
- : @interp_flat_type_eta' base_type_code var t T f x = f x.
- Proof using Type. t. Qed.
- (* Local *) Hint Rewrite @eq_interp_flat_type_eta'.
- Lemma interpf_exprf_eta {t e}
- : interpf (@interp_op) (exprf_eta (t:=t) e) = interpf (@interp_op) e.
- Proof using Type. t. Qed.
- (* Local *) Hint Rewrite @interpf_exprf_eta.
- Lemma interpf_exprf_eta' {t e}
- : interpf (@interp_op) (exprf_eta' (t:=t) e) = interpf (@interp_op) e.
- Proof using Type. t. Qed.
- (* Local *) Hint Rewrite @interpf_exprf_eta'.
- Lemma interp_expr_eta {t e}
- : forall x, interp (@interp_op) (expr_eta (t:=t) e) x = interp (@interp_op) e x.
- Proof using Type. t. Qed.
- Lemma interp_expr_eta' {t e}
- : forall x, interp (@interp_op) (expr_eta' (t:=t) e) x = interp (@interp_op) e x.
- Proof using Type. t. Qed.
- Lemma InterpExprEta {t e}
- : forall x, Interp (@interp_op) (ExprEta (t:=t) e) x = Interp (@interp_op) e x.
- Proof using Type. apply interp_expr_eta. Qed.
- Lemma InterpExprEta' {t e}
- : forall x, Interp (@interp_op) (ExprEta' (t:=t) e) x = Interp (@interp_op) e x.
- Proof using Type. apply interp_expr_eta'. Qed.
- Lemma InterpExprEta_arrow {s d e}
- : forall x, Interp (t:=Arrow s d) (@interp_op) (ExprEta (t:=Arrow s d) e) x = Interp (@interp_op) e x.
- Proof using Type. exact (@InterpExprEta (Arrow s d) e). Qed.
- Lemma InterpExprEta'_arrow {s d e}
- : forall x, Interp (t:=Arrow s d) (@interp_op) (ExprEta' (t:=Arrow s d) e) x = Interp (@interp_op) e x.
- Proof using Type. exact (@InterpExprEta' (Arrow s d) e). Qed.
-
- Lemma InterpExprEta_ind {t} (P : _ -> Prop) {e x}
- : P (Interp (@interp_op) e x) -> P (Interp (@interp_op) (ExprEta (t:=t) e) x).
- Proof using Type. rewrite InterpExprEta; exact id. Qed.
- Lemma InterpExprEta'_ind {t} (P : _ -> Prop) {e x}
- : P (Interp (@interp_op) e x) -> P (Interp (@interp_op) (ExprEta' (t:=t) e) x).
- Proof using Type. rewrite InterpExprEta'; exact id. Qed.
- Lemma InterpExprEta_arrow_ind {s d} (P : _ -> Prop) {e x}
- : P (Interp (@interp_op) e x) -> P (Interp (t:=Arrow s d) (@interp_op) (ExprEta (t:=Arrow s d) e) x).
- Proof using Type. rewrite InterpExprEta_arrow; exact id. Qed.
- Lemma InterpExprEta'_arrow_ind {s d} (P : _ -> Prop) {e x}
- : P (Interp (@interp_op) e x) -> P (Interp (t:=Arrow s d) (@interp_op) (ExprEta' (t:=Arrow s d) e) x).
- Proof using Type. rewrite InterpExprEta'_arrow; exact id. Qed.
-
- Lemma eq_interp_eta {t e}
- : forall x, interp_eta interp_op (t:=t) e x = interp interp_op e x.
- Proof using Type. apply eq_interp_flat_type_eta. Qed.
- Lemma eq_InterpEta {t e}
- : forall x, InterpEta interp_op (t:=t) e x = Interp interp_op e x.
- Proof using Type. apply eq_interp_eta. Qed.
-End language.
-
-Hint Rewrite @eq_interp_flat_type_eta @eq_interp_flat_type_eta' @interpf_exprf_eta @interpf_exprf_eta' @interp_expr_eta @interp_expr_eta' @InterpExprEta @InterpExprEta' @InterpExprEta_arrow @InterpExprEta'_arrow @eq_interp_eta @eq_InterpEta : reflective_interp.
diff --git a/src/Compilers/EtaWf.v b/src/Compilers/EtaWf.v
deleted file mode 100644
index 1134b5d05..000000000
--- a/src/Compilers/EtaWf.v
+++ /dev/null
@@ -1,122 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Eta.
-Require Import Crypto.Compilers.EtaInterp.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SplitInContext.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Local Notation exprf := (@exprf base_type_code op).
-
- Local Ltac t_step :=
- match goal with
- | _ => intro
- | _ => progress subst
- | _ => progress destruct_head' sig
- | _ => progress destruct_head' and
- | _ => progress simpl in *
- | _ => progress inversion_expr
- | _ => progress destruct_head' @expr
- | _ => progress invert_expr_step
- | [ |- iff _ _ ] => split
- | [ |- wf _ _ ] => constructor
- | _ => progress split_iff
- | _ => rewrite eq_interp_flat_type_eta_gen by assumption
- | [ H : _ |- _ ] => rewrite eq_interp_flat_type_eta_gen in H by assumption
- | [ H : context[interp_flat_type_eta_gen] |- _ ]
- => setoid_rewrite eq_interp_flat_type_eta_gen in H; [ | assumption.. ]
- | _ => progress break_match
- | [ H : wff _ _ _ |- _ ] => solve [ inversion H ]
- | [ |- wff _ _ _ ] => constructor
- | _ => solve [ auto | congruence | tauto ]
- end.
- Local Ltac t := repeat t_step.
-
- Local Hint Constructors wff.
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}.
- Section gen_flat_type.
- Context (eta : forall {A B}, A * B -> A * B)
- (eq_eta : forall A B x, @eta A B x = x).
- Section gen_type.
- Context (exprf_eta1 : forall {t} (e : exprf t), exprf t)
- (exprf_eta2 : forall {t} (e : exprf t), exprf t)
- (wff_exprf_eta : forall G t e1 e2, @wff _ _ var1 var2 G t e1 e2
- <-> @wff _ _ var1 var2 G t (@exprf_eta1 t e1) (@exprf_eta2 t e2)).
- Lemma wf_expr_eta_gen {t e1 e2}
- : wf (expr_eta_gen eta exprf_eta1 (t:=t) e1)
- (expr_eta_gen eta exprf_eta2 (t:=t) e2)
- <-> wf e1 e2.
- Proof using Type*. unfold expr_eta_gen; t; inversion_wf_step; t. Qed.
- End gen_type.
-
- Lemma wff_exprf_eta_gen {t e1 e2} G
- : wff G (exprf_eta_gen eta (t:=t) e1) (exprf_eta_gen eta (t:=t) e2)
- <-> @wff base_type_code op var1 var2 G t e1 e2.
- Proof using eq_eta.
- revert G; induction e1; first [ progress invert_expr | destruct e2 ];
- t; inversion_wf_step; t.
- Qed.
- End gen_flat_type.
-
- (* Local *) Hint Resolve -> wff_exprf_eta_gen.
- (* Local *) Hint Resolve <- wff_exprf_eta_gen.
-
- Lemma wff_exprf_eta {G t e1 e2}
- : wff G (exprf_eta (t:=t) e1) (exprf_eta (t:=t) e2)
- <-> @wff base_type_code op var1 var2 G t e1 e2.
- Proof using Type. setoid_rewrite wff_exprf_eta_gen; reflexivity. Qed.
- Lemma wff_exprf_eta' {G t e1 e2}
- : wff G (exprf_eta' (t:=t) e1) (exprf_eta' (t:=t) e2)
- <-> @wff base_type_code op var1 var2 G t e1 e2.
- Proof using Type. setoid_rewrite wff_exprf_eta_gen; intuition. Qed.
- Lemma wf_expr_eta {t e1 e2}
- : wf (expr_eta (t:=t) e1) (expr_eta (t:=t) e2)
- <-> @wf base_type_code op var1 var2 t e1 e2.
- Proof using Type.
- unfold expr_eta, exprf_eta.
- setoid_rewrite wf_expr_eta_gen; intuition (solve [ eapply wff_exprf_eta_gen; [ | eassumption ]; intuition ] || eauto).
- Qed.
- Lemma wf_expr_eta' {t e1 e2}
- : wf (expr_eta' (t:=t) e1) (expr_eta' (t:=t) e2)
- <-> @wf base_type_code op var1 var2 t e1 e2.
- Proof using Type.
- unfold expr_eta', exprf_eta'.
- setoid_rewrite wf_expr_eta_gen; intuition (solve [ eapply wff_exprf_eta_gen; [ | eassumption ]; intuition ] || eauto).
- Qed.
- End with_var.
-
- Lemma Wf_ExprEtaGen
- (eta : forall {A B}, A * B -> A * B)
- (eq_eta : forall A B x, @eta A B x = x)
- {t e}
- : Wf (ExprEtaGen (@eta) e) <-> @Wf base_type_code op t e.
- Proof using Type.
- split; intros H var1 var2; specialize (H var1 var2);
- revert H; eapply wf_expr_eta_gen; try eassumption; intros;
- symmetry; apply wff_exprf_eta_gen;
- auto.
- Qed.
- Lemma Wf_ExprEta_iff
- {t e}
- : Wf (ExprEta e) <-> @Wf base_type_code op t e.
- Proof using Type.
- unfold Wf; setoid_rewrite wf_expr_eta; reflexivity.
- Qed.
- Lemma Wf_ExprEta'_iff
- {t e}
- : Wf (ExprEta' e) <-> @Wf base_type_code op t e.
- Proof using Type.
- unfold Wf; setoid_rewrite wf_expr_eta'; reflexivity.
- Qed.
- Definition Wf_ExprEta {t e} : Wf e -> Wf (ExprEta e) := proj2 (@Wf_ExprEta_iff t e).
- Definition Wf_ExprEta' {t e} : Wf e -> Wf (ExprEta' e) := proj2 (@Wf_ExprEta'_iff t e).
-End language.
-
-Hint Resolve Wf_ExprEta Wf_ExprEta' : wf.
diff --git a/src/Compilers/ExprInversion.v b/src/Compilers/ExprInversion.v
deleted file mode 100644
index f3b469e83..000000000
--- a/src/Compilers/ExprInversion.v
+++ /dev/null
@@ -1,359 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.TypeInversion.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Notations.
-
-Section language.
- Context {base_type_code : Type}
- {interp_base_type : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type_gen := interp_flat_type.
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var : base_type_code -> Type}.
-
- Local Notation exprf := (@exprf base_type_code op var).
- Local Notation expr := (@expr base_type_code op var).
-
- Definition invert_Var {t} (e : exprf (Tbase t)) : option (var t)
- := match e in Syntax.exprf _ _ t'
- return option (var match t' with
- | Tbase t' => t'
- | _ => t
- end)
- with
- | Var _ v => Some v
- | _ => None
- end.
- Definition invert_Op {t} (e : exprf t) : option { t1 : flat_type & op t1 t * exprf t1 }%type
- := match e with Op _ _ opc args => Some (existT _ _ (opc, args)) | _ => None end.
- Definition invert_LetIn {A} (e : exprf A) : option { B : _ & exprf B * (Syntax.interp_flat_type var B -> exprf A) }%type
- := match e in Syntax.exprf _ _ t return option { B : _ & _ * (_ -> exprf t) }%type with
- | LetIn _ ex _ eC => Some (existT _ _ (ex, eC))
- | _ => None
- end.
- Definition invert_Pair {A B} (e : exprf (Prod A B)) : option (exprf A * exprf B)
- := match e in Syntax.exprf _ _ t
- return option match t with
- | Prod _ _ => _
- | _ => unit
- end with
- | Pair _ x _ y => Some (x, y)%core
- | _ => None
- end.
- Definition invert_Abs {T} (e : expr T) : interp_flat_type_gen var (domain T) -> exprf (codomain T)
- := match e with Abs _ _ f => f end.
-
- Section const.
- Context (invert_Const : forall s d, op s d -> exprf s -> option (interp_flat_type d)).
-
- Fixpoint lift_option {t} : interp_flat_type t -> interp_flat_type_gen (fun t => option (interp_base_type t)) t
- := match t with
- | Tbase T => fun x => Some x
- | Unit => fun _ => tt
- | Prod A B => fun (ab : interp_flat_type A * interp_flat_type B)
- => let '(a, b) := ab in
- (lift_option a, lift_option b)
- end.
-
- Fixpoint invert_PairsConst_gen {T} (e : exprf T)
- : option (interp_flat_type_gen (fun t => option (interp_base_type t)) T)
- := match e in Syntax.exprf _ _ t return option (interp_flat_type_gen (fun t => option (interp_base_type t)) t) with
- | TT => Some tt
- | Pair tx ex ty ey
- => match @invert_PairsConst_gen tx ex, @invert_PairsConst_gen ty ey with
- | Some x, Some y => Some (x, y)
- | Some _, None | None, Some _ | None, None => None
- end
- | Op s d opv args
- => option_map lift_option (invert_Const s d opv args)
- | Var _ _
- | LetIn _ _ _ _
- => None
- end.
- Fixpoint invert_PairsConst {T} (e : exprf T)
- : option (interp_flat_type T)
- := match e in Syntax.exprf _ _ t return option (interp_flat_type t) with
- | TT => Some tt
- | Pair tx ex ty ey
- => match @invert_PairsConst tx ex, @invert_PairsConst ty ey with
- | Some x, Some y => Some (x, y)
- | Some _, None | None, Some _ | None, None => None
- end
- | Op s d opv args
- => invert_Const s d opv args
- | Var _ _
- | LetIn _ _ _ _
- => None
- end.
- End const.
-
- Fixpoint invert_Pairs {T} (e : exprf T) : option (interp_flat_type_gen (fun ty => exprf (Tbase ty)) T)
- := match e in Syntax.exprf _ _ t
- return option (interp_flat_type_gen (fun ty => exprf (Tbase ty)) t)
- with
- | TT => Some tt
- | Var t _ as e => Some e
- | Pair tx ex ty ey
- => match @invert_Pairs tx ex, @invert_Pairs ty ey with
- | Some x, Some y => Some (x, y)
- | Some _, None | None, Some _ | None, None => None
- end
- | Op _ t _ _ as e
- | LetIn _ _ t _ as e
- => match t return exprf t -> option (interp_flat_type_gen _ t) with
- | Tbase _ => fun e => Some e
- | _ => fun _ => None
- end e
- end.
-
- Definition compose {A B C} (f : expr (B -> C)) (g : expr (A -> B))
- : expr (A -> C)
- := Abs (fun v => LetIn (invert_Abs g v)
- (invert_Abs f)).
-
- Definition exprf_code {t} (e : exprf t) : exprf t -> Prop
- := match e with
- | TT => fun e' => TT = e'
- | Var _ v => fun e' => invert_Var e' = Some v
- | Pair _ x _ y => fun e' => invert_Pair e' = Some (x, y)%core
- | Op _ _ opc args => fun e' => invert_Op e' = Some (existT _ _ (opc, args)%core)
- | LetIn _ ex _ eC => fun e' => invert_LetIn e' = Some (existT _ _ (ex, eC)%core)
- end.
-
- Definition expr_code {t} (e1 e2 : expr t) : Prop
- := invert_Abs e1 = invert_Abs e2.
-
- Definition exprf_encode {t} (x y : exprf t) : x = y -> exprf_code x y.
- Proof. intro p; destruct p, x; reflexivity. Defined.
- Definition expr_encode {t} (x y : expr t) : x = y -> expr_code x y.
- Proof. intro p; destruct p, x; reflexivity. Defined.
-
- Local Ltac t' :=
- repeat first [ intro
- | progress simpl in *
- | reflexivity
- | assumption
- | progress destruct_head False
- | progress subst
- | progress inversion_option
- | progress inversion_sigma
- | progress break_match ].
- Local Ltac t :=
- lazymatch goal with
- | [ |- _ = Some ?v -> ?e = _ ]
- => revert v;
- refine match e with
- | Var _ _ => _
- | _ => _
- end
- | [ |- _ = ?v -> ?e = _ ]
- => revert v;
- refine match e with
- | Abs _ _ _ => _
- end
- end;
- t'.
-
- Lemma invert_Var_Some {t e v}
- : @invert_Var t e = Some v -> e = Var v.
- Proof. t. Defined.
-
- Lemma invert_Op_Some {t e v}
- : @invert_Op t e = Some v -> e = Op (fst (projT2 v)) (snd (projT2 v)).
- Proof. t. Defined.
-
- Lemma invert_LetIn_Some {t e v}
- : @invert_LetIn t e = Some v -> e = LetIn (fst (projT2 v)) (snd (projT2 v)).
- Proof. t. Defined.
-
- Lemma invert_Pair_Some {A B e v}
- : @invert_Pair A B e = Some v -> e = Pair (fst v) (snd v).
- Proof. t. Defined.
-
- Lemma invert_Abs_Some {A B e v}
- : @invert_Abs (Arrow A B) e = v -> e = Abs v.
- Proof. t. Defined.
-
- Definition exprf_decode {t} (x y : exprf t) : exprf_code x y -> x = y.
- Proof.
- destruct x; simpl; trivial;
- intro H;
- first [ apply invert_Var_Some in H
- | apply invert_Op_Some in H
- | apply invert_LetIn_Some in H
- | apply invert_Pair_Some in H ];
- symmetry; assumption.
- Defined.
- Definition expr_decode {t} (x y : expr t) : expr_code x y -> x = y.
- Proof.
- destruct x; unfold expr_code; simpl.
- intro H; symmetry in H.
- apply invert_Abs_Some in H.
- symmetry; assumption.
- Defined.
- Definition path_exprf_rect {t} {x y : exprf t} (Q : x = y -> Type)
- (f : forall p, Q (exprf_decode x y p))
- : forall p, Q p.
- Proof. intro p; specialize (f (exprf_encode x y p)); destruct x, p; exact f. Defined.
- Definition path_expr_rect {t} {x y : expr t} (Q : x = y -> Type)
- (f : forall p, Q (expr_decode x y p))
- : forall p, Q p.
- Proof. intro p; specialize (f (expr_encode x y p)); destruct x, p; exact f. Defined.
- End with_var.
-
- Lemma interpf_invert_Abs interp_op {T e} x
- : Syntax.interpf interp_op (@invert_Abs interp_base_type T e x)
- = Syntax.interp interp_op e x.
- Proof using Type. destruct e; reflexivity. Qed.
-
- Lemma interpf_invert_PairsConst invert_Const interp_op {T} e v
- (Hinvert_Const
- : forall s d opc e v, invert_Const s d opc e = Some v
- -> interp_op s d opc (interpf interp_op e) = v)
- (H : invert_PairsConst (T:=T) invert_Const e = Some v)
- : Syntax.interpf interp_op e = v.
- Proof using Type.
- induction e;
- repeat first [ reflexivity
- | progress subst
- | solve [ auto ]
- | progress inversion_option
- | progress inversion_prod
- | progress simpl in *
- | progress break_innermost_match_hyps
- | apply (f_equal2 (@pair _ _)) ].
- Qed.
-
- Definition Compose {A B C} (f : Expr (B -> C)) (g : Expr (A -> B))
- : Expr (A -> C)
- := fun var => compose (f var) (g var).
-
- Lemma InterpCompose {A B C} interp_op f g
- : forall x, Interp interp_op (@Compose A B C f g) x
- = Interp interp_op f (Interp (interp_base_type:=interp_base_type) interp_op g x).
- Proof. reflexivity. Qed.
-End language.
-
-Global Arguments invert_Var {_ _ _ _} _.
-Global Arguments invert_Op {_ _ _ _} _.
-Global Arguments invert_LetIn {_ _ _ _} _.
-Global Arguments invert_Pair {_ _ _ _ _} _.
-Global Arguments invert_Pairs {_ _ _ _} _.
-Global Arguments invert_PairsConst {_ _ _ _} _ {T} _.
-Global Arguments invert_Abs {_ _ _ _} _ _.
-
-Hint Rewrite @InterpCompose : reflective_rewrite.
-
-Module Export Notations.
- Infix "∘" := Compose : expr_scope.
- Infix "∘f" := compose : expr_scope.
- Infix "∘ᶠ" := compose : expr_scope.
-End Notations.
-
-Ltac invert_one_expr e :=
- preinvert_one_type e;
- intros ? e;
- destruct e;
- try exact I.
-
-Ltac invert_expr_step :=
- match goal with
- | [ e : exprf _ _ (Tbase _) |- _ ] => invert_one_expr e
- | [ e : exprf _ _ (Prod _ _) |- _ ] => invert_one_expr e
- | [ e : exprf _ _ Unit |- _ ] => invert_one_expr e
- | [ e : expr _ _ (Arrow _ _) |- _ ] => invert_one_expr e
- end.
-
-Ltac invert_expr := repeat invert_expr_step.
-
-Ltac invert_match_expr_step :=
- match goal with
- | [ |- context[match ?e with TT => _ | _ => _ end] ]
- => invert_one_expr e
- | [ |- context[match ?e with Abs _ _ _ => _ end] ]
- => invert_one_expr e
- | [ H : context[match ?e with TT => _ | _ => _ end] |- _ ]
- => invert_one_expr e
- | [ H : context[match ?e with Abs _ _ _ => _ end] |- _ ]
- => invert_one_expr e
- end.
-
-Ltac invert_match_expr := repeat invert_match_expr_step.
-
-Ltac invert_expr_subst_step_helper guard_tac :=
- match goal with
- | [ H : invert_Var ?e = Some _ |- _ ] => guard_tac H; apply invert_Var_Some in H
- | [ H : invert_Op ?e = Some _ |- _ ] => guard_tac H; apply invert_Op_Some in H
- | [ H : invert_LetIn ?e = Some _ |- _ ] => guard_tac H; apply invert_LetIn_Some in H
- | [ H : invert_Pair ?e = Some _ |- _ ] => guard_tac H; apply invert_Pair_Some in H
- | [ e : expr _ _ _ |- _ ]
- => guard_tac e;
- let f := fresh e in
- let H := fresh in
- rename e into f;
- remember (invert_Abs f) as e eqn:H;
- symmetry in H;
- apply invert_Abs_Some in H;
- subst f
- | [ H : invert_Abs ?e = _ |- _ ] => guard_tac H; apply invert_Abs_Some in H
- end.
-Ltac invert_expr_subst_step :=
- first [ invert_expr_subst_step_helper ltac:(fun _ => idtac)
- | subst ].
-Ltac invert_expr_subst := repeat invert_expr_subst_step.
-
-Ltac induction_expr_in_using H rect :=
- induction H as [H] using (rect _ _ _);
- cbv [exprf_code expr_code invert_Var invert_LetIn invert_Pair invert_Op invert_Abs] in H;
- try lazymatch type of H with
- | Some _ = Some _ => apply option_leq_to_eq in H; unfold option_eq in H
- | Some _ = None => exfalso; clear -H; solve [ inversion H ]
- | None = Some _ => exfalso; clear -H; solve [ inversion H ]
- end;
- let H1 := fresh H in
- let H2 := fresh H in
- try lazymatch type of H with
- | existT _ _ _ = existT _ _ _ => induction_sigma_in_using H @path_sigT_rect
- end;
- try lazymatch type of H2 with
- | _ = (_, _)%core => induction_path_prod H2
- end.
-Ltac inversion_expr_step :=
- match goal with
- | [ H : _ = Var _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : _ = TT |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : _ = Op _ _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : _ = Pair _ _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : _ = LetIn _ _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : _ = Abs _ |- _ ]
- => induction_expr_in_using H @path_expr_rect
- | [ H : Var _ = _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : TT = _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : Op _ _ = _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : Pair _ _ = _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : LetIn _ _ = _ |- _ ]
- => induction_expr_in_using H @path_exprf_rect
- | [ H : Abs _ = _ |- _ ]
- => induction_expr_in_using H @path_expr_rect
- end.
-Ltac inversion_expr := repeat inversion_expr_step.
diff --git a/src/Compilers/FilterLive.v b/src/Compilers/FilterLive.v
deleted file mode 100644
index f54a8febb..000000000
--- a/src/Compilers/FilterLive.v
+++ /dev/null
@@ -1,70 +0,0 @@
-(** * Computes a list of live variables *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.CountLets.
-Require Import Crypto.Util.ListUtil.
-
-Local Notation eta x := (fst x, snd x).
-
-Local Open Scope ctype_scope.
-Section language.
- Context (base_type_code : Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (Name : Type)
- (dead_name : Name)
- (merge_names : Name -> Name -> Name)
- (* equations:
- - [merge_names x dead_name = merge_names dead_name x = x]
- - [merge_names x x = x] *).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation var := (fun t : base_type_code => list Name).
- Local Notation exprf := (@exprf base_type_code op var).
- Local Notation expr := (@expr base_type_code op var).
- Local Notation Expr := (@Expr base_type_code op var).
-
- Fixpoint merge_name_lists (ls1 ls2 : list Name) : list Name :=
- match ls1, ls2 with
- | cons x xs, cons y ys => cons (merge_names x y) (merge_name_lists xs ys)
- | ls1, nil => ls1
- | nil, ls2 => ls2
- end.
-
- Definition names_to_list {t} : interp_flat_type (fun _ : base_type_code => Name) t -> list Name
- := smart_interp_flat_map (g:=fun _ => list Name) (fun _ x => x :: nil)%list nil (fun _ _ x y => x ++ y)%list.
-
- Fixpoint filter_live_namesf (prefix remaining : list Name) {t} (e : exprf t) : list Name
- := match e with
- | TT => prefix
- | Var _ x => x
- | Op _ _ op args => @filter_live_namesf prefix remaining _ args
- | LetIn tx ex _ eC
- => let namesx := @filter_live_namesf prefix nil _ ex in
- let '(ns, remaining') := eta (split_names tx remaining) in
- match ns with
- | Some n =>
- @filter_live_namesf
- (prefix ++ List.repeat dead_name (count_pairs tx))%list remaining' _
- (eC (SmartValf (fun _ => list Name) (fun _ => namesx ++ names_to_list n)%list _))
- | None => nil
- end
- | Pair _ ex _ ey => merge_name_lists (@filter_live_namesf prefix remaining _ ex)
- (@filter_live_namesf prefix remaining _ ey)
- end.
-
- Definition filter_live_names (prefix remaining : list Name) {t} (e : expr t) : list Name
- := match e with
- | Abs src _ ef
- => let '(ns, remaining') := eta (split_names src remaining) in
- match ns with
- | Some n =>
- let prefix' := (prefix ++ names_to_list n)%list in
- filter_live_namesf
- prefix' remaining'
- (ef (SmartValf _ (fun _ => prefix') src))
- | None => nil
- end
- end.
-End language.
diff --git a/src/Compilers/FoldTypes.v b/src/Compilers/FoldTypes.v
deleted file mode 100644
index 0b923fcc9..000000000
--- a/src/Compilers/FoldTypes.v
+++ /dev/null
@@ -1,45 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.SmartMap.
-
-Section language.
- Context {base_type_code} {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Section generic_type.
- Context {A}
- (process : base_type_code -> A)
- (fold : A -> A -> A).
-
- Section with_var.
- Context {var : base_type_code -> Type}
- (init : A)
- (dummy : forall t, var t).
-
- Fixpoint fold_flat_type (t : flat_type base_type_code) : A
- := match t with
- | Tbase T => process T
- | Unit => init
- | Prod A B => fold (fold_flat_type A) (fold_flat_type B)
- end.
-
- Fixpoint type_foldf {t} (e : @exprf base_type_code op var t) : A
- := match e with
- | TT => init
- | Var t v => process t
- | Op t tR opc args
- => fold (@type_foldf t args) (fold_flat_type tR)
- | LetIn tx ex tC eC
- => fold (@type_foldf tx ex)
- (@type_foldf tC (eC (SmartValf _ dummy _)))
- | Pair tx ex ty ey
- => fold (@type_foldf tx ex) (@type_foldf ty ey)
- end.
-
- Definition type_fold {t} (e : @expr base_type_code op var t) : A
- := fold (fold_flat_type (domain t)) (type_foldf (invert_Abs e (SmartValf _ dummy _))).
- End with_var.
-
- Definition TypeFold (init : A) {t} (e : Expr base_type_code op t) : A
- := type_fold init (fun _ => tt) (e (fun _ => unit)).
- End generic_type.
-End language.
diff --git a/src/Compilers/GeneralizeVar.v b/src/Compilers/GeneralizeVar.v
deleted file mode 100644
index d73c3efd2..000000000
--- a/src/Compilers/GeneralizeVar.v
+++ /dev/null
@@ -1,39 +0,0 @@
-(** * Generalize [var] in [exprf] *)
-Require Import Coq.PArith.BinPos.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.InterpretToPHOAS.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Compilers.Syntax.
-
-(** N.B. This procedure only works when there are no nested lets,
- i.e., nothing like [let x := let y := z in w] in the PHOAS syntax
- tree. This is a limitation of [compile]. *)
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t)).
-
- Local Notation PContext var := (PositiveContext _ var _ base_type_code_bl_transparent).
-
- Definition GeneralizeVar {t} (e : @Syntax.expr base_type_code op (fun _ => FMapPositive.PositiveMap.key) t)
- : option (@Syntax.Expr base_type_code op (domain t -> codomain t))
- := let e := compile e (default_names_for (fun _ => 1%positive) e) in
- let e := match e with
- | Some e
- => match wf_unit (Context:=PContext _) empty e with
- | Some PointedProp.trivial => Some e
- | Some (PointedProp.inject _) => None
- | None => None
- end
- | None => None
- end in
- let e := option_map (InterpToPHOAS (Context:=fun var => PContext var) failb) e in
- e.
-End language.
diff --git a/src/Compilers/GeneralizeVarInterp.v b/src/Compilers/GeneralizeVarInterp.v
deleted file mode 100644
index aa7109ab1..000000000
--- a/src/Compilers/GeneralizeVarInterp.v
+++ /dev/null
@@ -1,79 +0,0 @@
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.InterpretToPHOASInterp.
-Require Import Crypto.Compilers.Named.CompileInterp.
-Require Import Crypto.Compilers.Named.CompileInterpSideConditions.
-Require Import Crypto.Compilers.Named.CompileWf.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Compilers.Named.PositiveContext.DefaultsProperties.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.GeneralizeVar.
-Require Import Crypto.Compilers.InterpSideConditions.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t))
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst).
-
- Local Notation GeneralizeVar
- := (@GeneralizeVar
- base_type_code op base_type_code_beq base_type_code_bl_transparent
- failb).
-
- Local Notation PositiveContextOk := (@PositiveContextOk base_type_code _ base_type_code_beq base_type_code_bl_transparent base_type_code_lb).
-
- Local Instance dec_base_type_code_eq : DecidableRel (@eq base_type_code).
- Proof.
- refine (fun x y => (if base_type_code_beq x y as b return base_type_code_beq x y = b -> Decidable (x = y)
- then fun pf => left (base_type_code_bl_transparent _ _ pf)
- else fun pf => right _) eq_refl).
- { clear -pf base_type_code_lb.
- let pf := pf in
- abstract (intro; erewrite base_type_code_lb in pf by eassumption; congruence). }
- Defined.
-
- Local Arguments Compile.compile : simpl never.
- Lemma interp_GeneralizeVar
- {t} (e1 e2 : expr base_type_code op t)
- (Hwf : wf e1 e2)
- e'
- (He' : GeneralizeVar e1 = Some e')
- : forall v, Interp interp_op e' v = interp interp_op e2 v.
- Proof using base_type_code_lb.
- unfold GeneralizeVar.GeneralizeVar, option_map in *.
- break_innermost_match_hyps; inversion_option; subst; intro.
- change (interp interp_op (?e ?var) ?v') with (Interp interp_op e v').
- unfold Interp, InterpretToPHOAS.Named.InterpToPHOAS, InterpretToPHOAS.Named.InterpToPHOAS_gen.
- match goal with |- ?L = ?R => cut (Some L = Some R); [ congruence | ] end.
- setoid_rewrite <- interp_interp_to_phoas.
- { erewrite (interp_compile (ContextOk:=PositiveContextOk)) with (e':=e2);
- [ reflexivity | auto | .. | eassumption ];
- auto using name_list_unique_default_names_for. }
- { eapply (wf_compile (ContextOk:=PositiveContextOk) (make_var':=fun _ => id)) with (e':= e2);
- [ auto | .. | eassumption ];
- auto using name_list_unique_default_names_for. }
- Qed.
-
- Lemma InterpGeneralizeVar
- {t} (e : Expr base_type_code op t)
- (Hwf : Wf e)
- e'
- (He' : GeneralizeVar (e _) = Some e')
- : forall v, Interp interp_op e' v = Interp interp_op e v.
- Proof using base_type_code_lb. eapply interp_GeneralizeVar; eauto. Qed.
-End language.
diff --git a/src/Compilers/GeneralizeVarWf.v b/src/Compilers/GeneralizeVarWf.v
deleted file mode 100644
index 9815d5092..000000000
--- a/src/Compilers/GeneralizeVarWf.v
+++ /dev/null
@@ -1,77 +0,0 @@
-Require Import Coq.PArith.BinPos.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.InterpretToPHOASWf.
-Require Import Crypto.Compilers.Named.WfFromUnit.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Compilers.Named.PositiveContext.DefaultsProperties.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.GeneralizeVar.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t))
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst).
-
- Local Notation GeneralizeVar
- := (@GeneralizeVar
- base_type_code op base_type_code_beq base_type_code_bl_transparent
- failb).
-
- Local Notation PositiveContextOk := (@PositiveContextOk base_type_code _ base_type_code_beq base_type_code_bl_transparent base_type_code_lb).
-
- Local Instance dec_base_type_code_eq : DecidableRel (@eq base_type_code).
- Proof using base_type_code_lb base_type_code_bl_transparent.
- refine (fun x y => (if base_type_code_beq x y as b return base_type_code_beq x y = b -> Decidable (x = y)
- then fun pf => left (base_type_code_bl_transparent _ _ pf)
- else fun pf => right _) eq_refl).
- { clear -pf base_type_code_lb.
- let pf := pf in
- abstract (intro; erewrite base_type_code_lb in pf by eassumption; congruence). }
- Defined.
-
- Local Arguments Compile.compile : simpl never.
- Lemma Wf_GeneralizeVar
- {t} (e1 : expr base_type_code op t)
- e'
- (He' : GeneralizeVar e1 = Some e')
- : Wf e'.
- Proof using base_type_code_lb.
- unfold GeneralizeVar, option_map in *.
- break_innermost_match_hyps; try congruence; intros v v'; [].
- inversion_option; subst.
- unfold InterpretToPHOAS.Named.InterpToPHOAS, InterpretToPHOAS.Named.InterpToPHOAS_gen.
- destruct t as [src dst].
- eapply (@wf_interp_to_phoas
- base_type_code op FMapPositive.PositiveMap.key _ _ _ _
- (PositiveContext base_type_code _ base_type_code_beq base_type_code_bl_transparent)
- (PositiveContext base_type_code _ base_type_code_beq base_type_code_bl_transparent)
- PositiveContextOk PositiveContextOk
- (failb _) (failb _) _ _);
- [ revert v | revert v' ];
- refine (_ : Wf.Named.Wf _ e);
- apply Wf_from_unit; auto using PositiveContextOk.
- Qed.
-
- Lemma Wf_GeneralizeVar_arrow
- {s d} (e : expr base_type_code op (Arrow s d))
- e'
- (He' : GeneralizeVar e = Some e')
- : Wf e'.
- Proof using base_type_code_lb. eapply Wf_GeneralizeVar; eassumption. Qed.
-End language.
-
-Hint Resolve Wf_GeneralizeVar Wf_GeneralizeVar_arrow : wf.
diff --git a/src/Compilers/InSet/Syntax.v b/src/Compilers/InSet/Syntax.v
deleted file mode 100644
index 4a3bb7a46..000000000
--- a/src/Compilers/InSet/Syntax.v
+++ /dev/null
@@ -1,84 +0,0 @@
-(** * PHOAS Representation of Gallina, in [Set] *)
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Compilers.Syntax.
-
-(** Universes are annoying. See, e.g.,
- [bug#5996](https://github.com/coq/coq/issues/5996), where using
- [pattern] and [constr:(...)] to replace [Set] with [Type] breaks
- things. Because we need to get reflective syntax by patterning [Z
- : Set], we make a version of [exprf] in [Set]. We build the
- minimial infrastructure needed to get this sort of expression into
- the [Type]-based one. *)
-
-Delimit Scope set_expr_scope with set_expr.
-Local Open Scope set_expr_scope.
-Section language.
- Context (base_type_code : Set).
-
- Local Notation flat_type := (flat_type base_type_code).
-
- Let Tbase' := @Tbase base_type_code.
- Local Coercion Tbase' : base_type_code >-> flat_type.
-
- Section interp.
- Context (interp_base_type : base_type_code -> Set).
- Fixpoint interp_flat_type (t : flat_type) :=
- match t with
- | Tbase t => interp_base_type t
- | Unit => unit
- | Prod x y => prod (interp_flat_type x) (interp_flat_type y)
- end.
- End interp.
-
- Section expr_param.
- Context (interp_base_type : base_type_code -> Set).
- Context (op : flat_type (* input tuple *) -> flat_type (* output type *) -> Set).
- Local Notation interp_flat_type_gen := interp_flat_type.
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Section expr.
- Context {var : base_type_code -> Set}.
-
- Inductive exprf : flat_type -> Set :=
- | TT : exprf Unit
- | Var {t} (v : var t) : exprf t
- | Op {t1 tR} (opc : op t1 tR) (args : exprf t1) : exprf tR
- | LetIn {tx} (ex : exprf tx) {tC} (eC : interp_flat_type_gen var tx -> exprf tC) : exprf tC
- | Pair {tx} (ex : exprf tx) {ty} (ey : exprf ty) : exprf (Prod tx ty).
- Bind Scope set_expr_scope with exprf.
- End expr.
-
- Section interp.
- Context (interp_op : forall src dst, op src dst -> interp_flat_type src -> interp_flat_type dst).
-
- Definition interpf_step
- (interpf : forall {t} (e : @exprf interp_flat_type t), interp_flat_type t)
- {t} (e : @exprf interp_flat_type t) : interp_flat_type t
- := match e in exprf t return interp_flat_type t with
- | TT => tt
- | Var _ x => x
- | Op _ _ op args => @interp_op _ _ op (@interpf _ args)
- | LetIn _ ex _ eC => dlet x := @interpf _ ex in @interpf _ (eC x)
- | Pair _ ex _ ey => (@interpf _ ex, @interpf _ ey)
- end.
-
- Fixpoint interpf {t} e
- := @interpf_step (@interpf) t e.
- End interp.
- End expr_param.
-End language.
-Global Arguments Var {_ _ _ _} _.
-Global Arguments TT {_ _ _}.
-Global Arguments Op {_ _ _ _ _} _ _.
-Global Arguments LetIn {_ _ _ _} _ {_} _.
-Global Arguments Pair {_ _ _ _} _ {_} _.
-Global Arguments Abs {_ _ _ _ _} _.
-Global Arguments interp_flat_type {_} _ _.
-Global Arguments interpf {_ _ _} interp_op {t} _.
-
-Module Export Notations.
- Notation "'slet' x .. y := A 'in' b" := (LetIn A%set_expr (fun x => .. (fun y => b%set_expr) .. )) : set_expr_scope.
- Notation "( x , y , .. , z )" := (Pair .. (Pair x%set_expr y%set_expr) .. z%set_expr) : set_expr_scope.
- Notation "( )" := TT : set_expr_scope.
- Notation "()" := TT : set_expr_scope.
-End Notations.
diff --git a/src/Compilers/InSet/Typeify.v b/src/Compilers/InSet/Typeify.v
deleted file mode 100644
index d2fbad987..000000000
--- a/src/Compilers/InSet/Typeify.v
+++ /dev/null
@@ -1,58 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InSet.Syntax.
-
-Section language.
- Context {base_type_code : Set}
- {op : flat_type base_type_code -> flat_type base_type_code -> Set}
- {var : base_type_code -> Set}.
-
- Fixpoint typeify_interp_flat_type {t}
- : InSet.Syntax.interp_flat_type var t -> Compilers.Syntax.interp_flat_type var t
- := match t with
- | Tbase T => fun v => v
- | Unit => fun v => v
- | Prod A B => fun ab : InSet.Syntax.interp_flat_type _ A * InSet.Syntax.interp_flat_type _ B
- => (@typeify_interp_flat_type _ (fst ab),
- @typeify_interp_flat_type _ (snd ab))
- end.
- Fixpoint untypeify_interp_flat_type {t}
- : Compilers.Syntax.interp_flat_type var t -> InSet.Syntax.interp_flat_type var t
- := match t with
- | Tbase T => fun v => v
- | Unit => fun v => v
- | Prod A B => fun ab : Compilers.Syntax.interp_flat_type _ A * Compilers.Syntax.interp_flat_type _ B
- => (@untypeify_interp_flat_type _ (fst ab),
- @untypeify_interp_flat_type _ (snd ab))
- end.
-
- Fixpoint typeify {t} (e : @InSet.Syntax.exprf base_type_code op var t)
- : @Compilers.Syntax.exprf base_type_code op var t
- := match e with
- | TT => Compilers.Syntax.TT
- | Var t v => Compilers.Syntax.Var v
- | Op t1 tR opc args => Compilers.Syntax.Op opc (@typeify _ args)
- | LetIn tx ex tC eC
- => Compilers.Syntax.LetIn
- (@typeify _ ex)
- (fun x => @typeify _ (eC (untypeify_interp_flat_type x)))
- | Pair tx ex ty ey => Compilers.Syntax.Pair
- (@typeify _ ex)
- (@typeify _ ey)
- end.
-
- Fixpoint untypeify {t} (e : @Compilers.Syntax.exprf base_type_code op var t)
- : @InSet.Syntax.exprf base_type_code op var t
- := match e with
- | Compilers.Syntax.TT => TT
- | Compilers.Syntax.Var t v => Var v
- | Compilers.Syntax.Op t1 tR opc args => Op opc (@untypeify _ args)
- | Compilers.Syntax.LetIn tx ex tC eC
- => LetIn
- (@untypeify _ ex)
- (fun x => @untypeify _ (eC (typeify_interp_flat_type x)))
- | Compilers.Syntax.Pair tx ex ty ey
- => Pair
- (@untypeify _ ex)
- (@untypeify _ ey)
- end.
-End language.
diff --git a/src/Compilers/InSet/TypeifyInterp.v b/src/Compilers/InSet/TypeifyInterp.v
deleted file mode 100644
index 77790eec8..000000000
--- a/src/Compilers/InSet/TypeifyInterp.v
+++ /dev/null
@@ -1,114 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InSet.Syntax.
-Require Import Crypto.Compilers.InSet.Typeify.
-
-
-Local Ltac t :=
- repeat first [ reflexivity
- | progress simpl in *
- | apply (f_equal2 (@pair _ _))
- | solve [ auto ]
- | progress cbv [LetIn.Let_In]
- | progress autorewrite with core
- | progress intros
- | match goal with
- | [ v : _ * _ |- _ ] => destruct v
- | [ H : _ |- _ ] => rewrite H
- end ].
-
-Section language.
- Context {base_type_code : Set}
- {op : flat_type base_type_code -> flat_type base_type_code -> Set}
- {interp_base_type : base_type_code -> Set}.
-
- Lemma typeify_untypeify_interp_flat_type {t} v
- : @typeify_interp_flat_type base_type_code interp_base_type t (untypeify_interp_flat_type v) = v.
- Proof using Type. induction t; t. Qed.
- Lemma untypeify_typeify_interp_flat_type {t} v
- : @untypeify_interp_flat_type base_type_code interp_base_type t (typeify_interp_flat_type v) = v.
- Proof using Type. induction t; t. Qed.
- Hint Rewrite @typeify_untypeify_interp_flat_type @untypeify_typeify_interp_flat_type.
-
- Section gen.
- Context {interp_op : forall s d, op s d -> Compilers.Syntax.interp_flat_type interp_base_type s -> Compilers.Syntax.interp_flat_type interp_base_type d}
- {interp_op' : forall s d, op s d -> InSet.Syntax.interp_flat_type interp_base_type s -> InSet.Syntax.interp_flat_type interp_base_type d}
- (untypeify_interp_op
- : forall s d opc args,
- untypeify_interp_flat_type (interp_op s d opc args)
- = interp_op' s d opc (untypeify_interp_flat_type args))
- (typeify_interp_op
- : forall s d opc args,
- typeify_interp_flat_type (interp_op' s d opc args)
- = interp_op s d opc (typeify_interp_flat_type args)).
-
- Local Notation interpf := (Compilers.Syntax.interpf interp_op).
- Local Notation interpf' := (InSet.Syntax.interpf interp_op').
- Local Notation typeify := (typeify (var:=interp_base_type)).
-
- Lemma interpf_typeify {t} e
- : interpf (typeify e) = typeify_interp_flat_type (t:=t) (interpf' e).
- Proof using typeify_interp_op. induction e; t. Qed.
- Lemma interpf_untypeify {t} e
- : interpf' (untypeify e) = untypeify_interp_flat_type (t:=t) (interpf e).
- Proof using untypeify_interp_op. induction e; t. Qed.
-
- Lemma interpf_typeify_untypeify {t} e
- : interpf (typeify (untypeify (t:=t) e)) = interpf e.
- Proof using Type. induction e; t. Qed.
- Lemma interpf_untypeify_typeify {t} e
- : interpf' (untypeify (t:=t) (typeify e)) = interpf' e.
- Proof using Type. induction e; t. Qed.
- End gen.
-End language.
-
-Module Compilers.
- Import Compilers.Syntax.
- Section language.
- Context {base_type_code : Set}
- {op : flat_type base_type_code -> flat_type base_type_code -> Set}
- {interp_base_type : base_type_code -> Set}
- {interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}.
-
- Hint Rewrite @typeify_untypeify_interp_flat_type @untypeify_typeify_interp_flat_type.
-
- Local Notation interp_op' :=
- (fun s d opc args => untypeify_interp_flat_type (interp_op s d opc (typeify_interp_flat_type args))).
-
- Local Notation interpf := (Compilers.Syntax.interpf interp_op).
- Local Notation interpf' := (InSet.Syntax.interpf interp_op').
- Local Notation typeify := (typeify (var:=interp_base_type)).
-
- Lemma interpf_typeify {t} e
- : interpf (typeify e) = typeify_interp_flat_type (t:=t) (interpf' e).
- Proof using Type. apply interpf_typeify; t. Qed.
- Lemma interpf_untypeify {t} e
- : interpf' (untypeify e) = untypeify_interp_flat_type (t:=t) (interpf e).
- Proof using Type. apply interpf_untypeify; t. Qed.
- End language.
-End Compilers.
-
-Module InSet.
- Import InSet.Syntax.
- Section language.
- Context {base_type_code : Set}
- {op : flat_type base_type_code -> flat_type base_type_code -> Set}
- {interp_base_type : base_type_code -> Set}
- {interp_op' : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}.
-
- Hint Rewrite @typeify_untypeify_interp_flat_type @untypeify_typeify_interp_flat_type.
-
- Local Notation interp_op :=
- (fun s d opc args => typeify_interp_flat_type (interp_op' s d opc (untypeify_interp_flat_type args))).
-
- Local Notation interpf := (Compilers.Syntax.interpf interp_op).
- Local Notation interpf' := (InSet.Syntax.interpf interp_op').
- Local Notation typeify := (typeify (var:=interp_base_type)).
-
- Lemma interpf_typeify {t} e
- : interpf (typeify e) = typeify_interp_flat_type (t:=t) (interpf' e).
- Proof using Type. apply interpf_typeify; t. Qed.
- Lemma interpf_untypeify {t} e
- : interpf' (untypeify e) = untypeify_interp_flat_type (t:=t) (interpf e).
- Proof using Type. apply interpf_untypeify; t. Qed.
- End language.
-End InSet.
diff --git a/src/Compilers/Inline.v b/src/Compilers/Inline.v
deleted file mode 100644
index bd39b48a0..000000000
--- a/src/Compilers/Inline.v
+++ /dev/null
@@ -1,104 +0,0 @@
-(** * Inline: Remove some [Let] expressions *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Let Tbase := @Tbase base_type_code.
- Local Coercion Tbase : base_type_code >-> Syntax.flat_type.
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var : base_type_code -> Type}.
-
- Inductive inline_directive : flat_type -> Type :=
- | default_inline {t} (e : @exprf var t) : inline_directive t
- | inline {t} (e : interp_flat_type (fun t => @exprf var (Tbase t)) t) : inline_directive t
- | no_inline {t} (e : @exprf var t) : inline_directive t
- | partial_inline
- {tx tC}
- (ex : @exprf var tx)
- (eC : interp_flat_type var tx -> interp_flat_type (fun t => @exprf var (Tbase t)) tC)
- : inline_directive tC.
-
- Definition exprf_of_inline_directive {t} (v : inline_directive t) : @exprf var t
- := match v with
- | default_inline t e => e
- | inline t e => SmartPairf e
- | no_inline t e => e
- | partial_inline _ _ ex eC
- => LetIn ex (fun x => SmartPairf (eC x))
- end.
-
- Context (postprocess : forall {t}, @exprf var t -> inline_directive t).
-
- Fixpoint inline_const_genf {t} (e : @exprf (@exprf var) t) : @exprf var t
- := match e in Syntax.exprf _ _ t return @exprf var t with
- | LetIn tx ex tC eC
- => match postprocess _ (@inline_const_genf _ ex) in inline_directive t' return (interp_flat_type _ t' -> @exprf var tC) -> @exprf var tC with
- | default_inline _ ex
- => match ex in Syntax.exprf _ _ t' return (interp_flat_type _ t' -> @exprf var tC) -> @exprf var tC with
- | TT => fun eC => eC tt
- | Var _ x => fun eC => eC (Var x)
- | ex => fun eC => LetIn ex (fun x => eC (SmartVarVarf x))
- end
- | no_inline _ ex
- => fun eC => LetIn ex (fun x => eC (SmartVarVarf x))
- | inline _ ex => fun eC => eC ex
- | partial_inline _ _ ex eC'
- => fun eC => LetIn ex (fun x => eC (eC' x))
- end (fun x => @inline_const_genf _ (eC x))
- | Var _ x => x
- | TT => TT
- | Pair _ ex _ ey => Pair (@inline_const_genf _ ex) (@inline_const_genf _ ey)
- | Op _ _ op args => Op op (@inline_const_genf _ args)
- end.
-
- Definition inline_const_gen {t} (e : @expr (@exprf var) t) : @expr var t
- := match e in Syntax.expr _ _ t return @expr var t with
- | Abs _ _ f => Abs (fun x => inline_const_genf (f (SmartVarVarf x)))
- end.
-
- Section with_is_const.
- Context (is_const : forall s d, op s d -> bool).
-
- Definition postprocess_for_const (t : flat_type) (v : @exprf var t) : inline_directive t
- := if match v with Op _ _ op _ => @is_const _ _ op | _ => false end
- then match t return @exprf _ t -> inline_directive t with
- | Syntax.Tbase _ => @inline (Tbase _)
- | _ => @default_inline _
- end v
- else default_inline v.
- End with_is_const.
- End with_var.
-
- Definition inline_constf is_const {var t}
- := @inline_const_genf var (postprocess_for_const is_const) t.
- Definition inline_const is_const {var t}
- := @inline_const_gen var (postprocess_for_const is_const) t.
-
- Definition InlineConstGen (postprocess : forall var t, @exprf var t -> @inline_directive var t)
- {t} (e : Expr t) : Expr t
- := fun var => inline_const_gen (postprocess _) (e _).
- Definition InlineConst is_const {t}
- := @InlineConstGen (fun var => postprocess_for_const is_const) t.
-End language.
-
-Global Arguments inline_directive {_} _ _ _, {_ _ _} _.
-Global Arguments no_inline {_ _ _ _} _.
-Global Arguments inline {_ _ _ _} _.
-Global Arguments default_inline {_ _ _ _} _.
-Global Arguments partial_inline {_ _ _ _ _} ex eC.
-Global Arguments inline_const_genf {_ _ _} postprocess {_} _.
-Global Arguments inline_const_gen {_ _ _} postprocess {_} _.
-Global Arguments InlineConstGen {_ _} postprocess {_} _ var.
-Global Arguments inline_constf {_ _} is_const {_ t} _.
-Global Arguments inline_const {_ _} is_const {_ t} _.
-Global Arguments InlineConst {_ _} is_const {_} _ var.
diff --git a/src/Compilers/InlineConstAndOp.v b/src/Compilers/InlineConstAndOp.v
deleted file mode 100644
index 842fc628a..000000000
--- a/src/Compilers/InlineConstAndOp.v
+++ /dev/null
@@ -1,149 +0,0 @@
-(** * Inline: Remove some [Let] expressions, inline constants, interpret constant operations *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Inline.
-Require Import Crypto.Compilers.ExprInversion.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var : base_type_code -> Type}
- (invert_Const : forall s d, op s d -> @exprf var s -> option (interp_flat_type interp_base_type d))
- (Const : forall t, interp_base_type t -> @exprf var (Tbase t)).
-
- Local Notation invert_PairsConst' T
- := (@invert_PairsConst base_type_code interp_base_type op var invert_Const T).
- Local Notation invert_PairsConst
- := (invert_PairsConst' _).
-
- Fixpoint postprocess_for_const_and_op {t} (e : exprf t)
- : inline_directive (var:=var) (op:=op) t
- := match e with
- | TT => inline (t:=Unit) tt
- | Var t v => inline (t:=Tbase _) (Var v)
- | Op t1 tR opc args as e
- => match invert_PairsConst args with
- | Some args
- => inline (SmartVarfMap Const (interp_op _ _ opc args))
- | None
- => no_inline e
- end
- | LetIn _ _ _ _ as e
- => no_inline e
- | Pair tx ex ty ey as e
- => match @postprocess_for_const_and_op _ ex in inline_directive tx, @postprocess_for_const_and_op _ ey in inline_directive ty
- return inline_directive (Prod tx ty) -> inline_directive (Prod tx ty)
- with
- | inline tx ex, inline ty ey
- => fun _ => inline (t:=Prod tx ty) (ex, ey)
- | partial_inline tx tC ex eC, partial_inline ty tC' ey eC'
- => fun _ => partial_inline
- (tC:=Prod _ _)
- (Pair ex ey)
- (fun xy : interp_flat_type _ _ * interp_flat_type _ _
- => (eC (fst xy), eC' (snd xy)))
- | no_inline _ ex, no_inline _ ey
- => fun _ => no_inline (Pair ex ey)
- | no_inline tx ex, inline ty ey
- => fun _ => partial_inline
- (tC:=Prod _ _)
- ex (fun x => (SmartVarVarf x, ey))
- | inline tx ex, no_inline ty ey
- => fun _ => partial_inline
- (tC:=Prod _ _)
- ey (fun y => (ex, SmartVarVarf y))
- | partial_inline tx tC ex eC, inline ty ey
- => fun _ => partial_inline
- (tC:=Prod _ _)
- ex (fun x => (eC x, ey))
- | inline tx ex, partial_inline ty tC ey eC
- => fun _ => partial_inline
- (tC:=Prod _ _)
- ey (fun y => (ex, eC y))
- | partial_inline tx tC ex eC, no_inline ty ey
- => fun _ => partial_inline
- (tC:=Prod _ _)
- (Pair ex ey)
- (fun xy : interp_flat_type _ _ * interp_flat_type _ _
- => (eC (fst xy), SmartVarVarf (snd xy)))
- | no_inline tx ex, partial_inline ty tC ey eC
- => fun _ => partial_inline
- (tC:=Prod _ _)
- (Pair ex ey)
- (fun xy : interp_flat_type _ _ * interp_flat_type _ _
- => (SmartVarVarf (fst xy), eC (snd xy)))
- | default_inline _ ex, default_inline _ ey
- => fun d => d
- | default_inline _ _, _
- | _, default_inline _ _
- => fun d => d
- end (default_inline e)
- end.
-
- Definition inline_const_and_op_genf {t}
- := @inline_const_genf base_type_code op var (@postprocess_for_const_and_op) t.
-
- Definition inline_const_and_op_gen {t}
- := @inline_const_gen base_type_code op var (@postprocess_for_const_and_op) t.
- End with_var.
-
- Section const_unit.
- Context {var : base_type_code -> Type}
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t)).
-
- Definition invert_ConstUnit' {s d} : op s d -> option (interp_flat_type interp_base_type d)
- := match s with
- | Unit
- => fun opv => Some (interp_op _ _ opv tt)
- | _ => fun _ => None
- end.
- Definition invert_ConstUnit {s d} (opv : op s d) (e : @exprf var s)
- : option (interp_flat_type interp_base_type d)
- := invert_ConstUnit' opv.
-
- Definition Const {t} v := Op (var:=var) (OpConst t v) TT.
-
- Definition inline_const_and_opf {t}
- := @inline_const_and_op_genf var (@invert_ConstUnit) (@Const) t.
- Definition inline_const_and_op {t}
- := @inline_const_and_op_gen var (@invert_ConstUnit) (@Const) t.
- End const_unit.
-
- Definition InlineConstAndOpGen
- (invert_Const : forall var s d, op s d -> @exprf var s -> option (interp_flat_type interp_base_type d))
- (Const : forall var t, interp_base_type t -> @exprf var (Tbase t))
- {t} (e : Expr t)
- : Expr t
- := @InlineConstGen
- base_type_code op
- (fun var => @postprocess_for_const_and_op var (invert_Const _) (Const _))
- t
- e.
-
- Definition InlineConstAndOp
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- {t} (e : Expr t)
- : Expr t
- := @InlineConstGen
- base_type_code op
- (fun var => @postprocess_for_const_and_op var (@invert_ConstUnit _) (@Const _ OpConst))
- t
- e.
-End language.
-
-Global Arguments inline_const_and_op_genf {_ _ _} interp_op {var} invert_Const Const {t} _ : assert.
-Global Arguments inline_const_and_op_gen {_ _ _} interp_op {var} invert_Const Const {t} _ : assert.
-Global Arguments inline_const_and_opf {_ _ _} interp_op {var} OpConst {t} _ : assert.
-Global Arguments inline_const_and_op {_ _ _} interp_op {var} OpConst {t} _ : assert.
-Global Arguments InlineConstAndOpGen {_ _ _} interp_op invert_Const Const {t} _ var : assert.
-Global Arguments InlineConstAndOp {_ _ _} interp_op OpConst {t} _ var : assert.
diff --git a/src/Compilers/InlineConstAndOpByRewrite.v b/src/Compilers/InlineConstAndOpByRewrite.v
deleted file mode 100644
index d1d610f8b..000000000
--- a/src/Compilers/InlineConstAndOpByRewrite.v
+++ /dev/null
@@ -1,93 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Rewriter.
-
-Module Export Rewrite.
- Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d).
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var : base_type_code -> Type}
- (invert_Const : forall s d, op s d -> @exprf var s -> option (interp_flat_type interp_base_type d))
- (Const : forall t, interp_base_type t -> @exprf var (Tbase t)).
-
- Local Notation invert_PairsConst' T
- := (@invert_PairsConst base_type_code interp_base_type op var invert_Const T).
- Local Notation invert_PairsConst
- := (invert_PairsConst' _).
-
- Definition rewrite_for_const_and_op src dst (opc : op src dst) (args : @exprf var src)
- : @exprf var dst
- := match invert_PairsConst args with
- | Some argsv
- => SmartPairf (SmartVarfMap Const (interp_op _ _ opc argsv))
- | None => Op opc args
- end.
-
- Definition inline_const_and_op_genf {t} (e : @exprf var t) : @exprf var t
- := @rewrite_opf base_type_code op var rewrite_for_const_and_op t e.
- Definition inline_const_and_op_gen {t} (e : @expr var t) : @expr var t
- := @rewrite_op base_type_code op var rewrite_for_const_and_op t e.
- End with_var.
-
-
- Section const_unit.
- Context {var : base_type_code -> Type}
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t)).
-
- Definition invert_ConstUnit' {s d} : op s d -> option (interp_flat_type interp_base_type d)
- := match s with
- | Unit
- => fun opv => Some (interp_op _ _ opv tt)
- | _ => fun _ => None
- end.
- Definition invert_ConstUnit {s d} (opv : op s d) (e : @exprf var s)
- : option (interp_flat_type interp_base_type d)
- := invert_ConstUnit' opv.
-
- Definition Const {t} v := Op (var:=var) (OpConst t v) TT.
-
- Definition inline_const_and_opf {t}
- := @inline_const_and_op_genf var (@invert_ConstUnit) (@Const) t.
- Definition inline_const_and_op {t}
- := @inline_const_and_op_gen var (@invert_ConstUnit) (@Const) t.
- End const_unit.
-
- Definition InlineConstAndOpGen
- (invert_Const : forall var s d, op s d -> @exprf var s -> option (interp_flat_type interp_base_type d))
- (Const : forall var t, interp_base_type t -> @exprf var (Tbase t))
- {t} (e : Expr t)
- : Expr t
- := @RewriteOp
- base_type_code op
- (fun var => @rewrite_for_const_and_op var (invert_Const _) (Const _))
- t
- e.
-
- Definition InlineConstAndOp
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- {t} (e : Expr t)
- : Expr t
- := @RewriteOp
- base_type_code op
- (fun var => @rewrite_for_const_and_op var (@invert_ConstUnit _) (@Const _ OpConst))
- t
- e.
- End language.
-
- Global Arguments inline_const_and_op_genf {_ _ _} interp_op {var} invert_Const Const {t} _ : assert.
- Global Arguments inline_const_and_op_gen {_ _ _} interp_op {var} invert_Const Const {t} _ : assert.
- Global Arguments inline_const_and_opf {_ _ _} interp_op {var} OpConst {t} _ : assert.
- Global Arguments inline_const_and_op {_ _ _} interp_op {var} OpConst {t} _ : assert.
- Global Arguments InlineConstAndOpGen {_ _ _} interp_op invert_Const Const {t} _ var : assert.
- Global Arguments InlineConstAndOp {_ _ _} interp_op OpConst {t} _ var : assert.
-End Rewrite.
diff --git a/src/Compilers/InlineConstAndOpByRewriteInterp.v b/src/Compilers/InlineConstAndOpByRewriteInterp.v
deleted file mode 100644
index 33e525ed4..000000000
--- a/src/Compilers/InlineConstAndOpByRewriteInterp.v
+++ /dev/null
@@ -1,136 +0,0 @@
-(** * Inline: Remove some [Let] expressions, inline constants, interpret constant operations *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Rewriter.
-Require Import Crypto.Compilers.RewriterInterp.
-Require Import Crypto.Compilers.InlineConstAndOpByRewrite.
-Require Import Crypto.Util.Sigma Crypto.Util.Prod Crypto.Util.Option.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Module Export Rewrite.
- Local Open Scope ctype_scope.
- Section language.
- Context (base_type_code : Type)
- (interp_base_type : base_type_code -> Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_invert.
- Context (invert_Const : forall s d, op s d -> @exprf interp_base_type s -> option (interp_flat_type d))
- (Const : forall t, interp_base_type t -> @exprf interp_base_type (Tbase t))
- (Hinvert_Const
- : forall s d opc e v, invert_Const s d opc e = Some v
- -> interp_op s d opc (interpf interp_op e) = v)
- (interpf_Const : forall t v, interpf interp_op (Const t v) = v).
-
- Lemma interpf_rewrite_for_const_and_op s d opc args
- : interpf interp_op (rewrite_for_const_and_op interp_op invert_Const Const s d opc args)
- = interp_op _ _ opc (interpf interp_op args).
- Proof using Hinvert_Const interpf_Const.
- cbv [rewrite_for_const_and_op].
- repeat first [ break_innermost_match_step
- | reflexivity
- | progress cbn [interpf exprf_of_inline_directive interpf_step LetIn.Let_In SmartVarVarf fst snd] in *
- | rewrite interpf_SmartPairf
- | rewrite SmartVarfMap_compose
- | rewrite SmartVarfMap_id
- | setoid_rewrite interpf_Const
- | erewrite interpf_invert_PairsConst by eassumption ].
- Qed.
-
- Lemma interpf_inline_const_and_op_genf
- {t} e
- : interpf interp_op (inline_const_and_op_genf (t:=t) interp_op invert_Const Const e)
- = interpf interp_op e.
- Proof.
- unfold inline_const_and_op_genf;
- eapply interpf_rewrite_opf; eauto using interpf_rewrite_for_const_and_op.
- Qed.
-
- Lemma interpf_inline_const_and_op_gen
- {t} e
- : forall x,
- interp interp_op (inline_const_and_op_gen (t:=t) interp_op invert_Const Const e) x
- = interp interp_op e x.
- Proof.
- unfold inline_const_and_op_gen;
- eapply interp_rewrite_op; eauto using interpf_rewrite_for_const_and_op.
- Qed.
- End with_invert.
-
- Section const_unit.
- Context (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- (interp_op_OpConst : forall t v, interp_op _ _ (OpConst t v) tt = v).
-
- Lemma interpf_invert_ConstUnit s d opc e v
- (H : invert_ConstUnit interp_op opc e = Some v)
- : interp_op s d opc (interpf interp_op e) = v.
- Proof using Type.
- destruct s; simpl in *; inversion_option; subst.
- edestruct interpf; reflexivity.
- Qed.
-
- Lemma interpf_Const t v
- : interpf interp_op (Const OpConst (t:=t) v) = v.
- Proof using interp_op_OpConst.
- apply interp_op_OpConst.
- Qed.
-
- Lemma interpf_inline_const_and_opf
- {t} e
- : interpf interp_op (inline_const_and_opf (t:=t) interp_op OpConst e)
- = interpf interp_op e.
- Proof.
- unfold inline_const_and_opf;
- eapply interpf_rewrite_opf; eauto using interpf_rewrite_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
-
- Lemma interpf_inline_const_and_op
- {t} e
- : forall x,
- interp interp_op (inline_const_and_op (t:=t) interp_op OpConst e) x
- = interp interp_op e x.
- Proof.
- unfold inline_const_and_op;
- eapply interp_rewrite_op; eauto using interpf_rewrite_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
- End const_unit.
-
- Lemma InterpInlineConstAndOpGen
- (invert_Const : forall var s d, op s d -> @exprf var s -> option (interp_flat_type d))
- (Const : forall var t, interp_base_type t -> @exprf var (Tbase t))
-
- {t} (e : Expr t)
- (Hinvert_Const
- : forall s d opc e v,
- invert_Const _ s d opc e = Some v
- -> interp_op s d opc (interpf interp_op e) = v)
- (interpf_Const : forall t v, interpf interp_op (Const _ t v) = v)
- : forall x, Interp interp_op (InlineConstAndOpGen interp_op invert_Const Const e) x = Interp interp_op e x.
- Proof using Type.
- eapply InterpRewriteOp;
- eauto using interpf_rewrite_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
-
- Lemma InterpInlineConstAndOp
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- {t} (e : Expr t)
- (interp_op_OpConst : forall t v, interp_op _ _ (OpConst t v) tt = v)
- : forall x, Interp interp_op (InlineConstAndOp interp_op OpConst e) x = Interp interp_op e x.
- Proof using Type.
- eapply InterpRewriteOp;
- eauto using interpf_rewrite_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
- End language.
-
- Hint Rewrite @InterpInlineConstAndOp @InterpInlineConstAndOpGen using assumption : reflective_interp.
-End Rewrite.
diff --git a/src/Compilers/InlineConstAndOpByRewriteWf.v b/src/Compilers/InlineConstAndOpByRewriteWf.v
deleted file mode 100644
index c1d06be99..000000000
--- a/src/Compilers/InlineConstAndOpByRewriteWf.v
+++ /dev/null
@@ -1,171 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Rewriter.
-Require Import Crypto.Compilers.RewriterWf.
-Require Import Crypto.Compilers.InlineConstAndOpByRewrite.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Option.
-
-Module Export Rewrite.
- Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation Tbase := (@Tbase base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
-
- Local Hint Constructors Wf.wff.
-
- Section with_var.
- Context {interp_op1 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {interp_op2 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {var1 var2 : base_type_code -> Type}
- {invert_Const1 : forall s d, op s d -> @exprf var1 s -> option (interp_flat_type interp_base_type d)}
- {invert_Const2 : forall s d, op s d -> @exprf var2 s -> option (interp_flat_type interp_base_type d)}
- {Const1 : forall t, interp_base_type t -> @exprf var1 (Tbase t)}
- {Const2 : forall t, interp_base_type t -> @exprf var2 (Tbase t)}
- (Hinterp_op : forall s d opv args, interp_op1 s d opv args = interp_op2 s d opv args)
- (Hinvert_Const : forall s d opv G e1 e2, wff G e1 e2 -> invert_Const1 s d opv e1 = invert_Const2 s d opv e2)
- (HConst : forall t v G, wff G (Const1 t v) (Const2 t v)).
-
-
- Local Ltac t_fin :=
- repeat first [ intro
- | exfalso; assumption
- | exact I
- | progress destruct_head'_prod
- | progress destruct_head'_and
- | progress destruct_head'_sig
- | progress subst
- | rewrite Hinterp_op
- | rewrite SmartPairf_Pair
- | tauto
- | solve [ auto with nocore
- | apply wff_SmartPairf_SmartVarfMap_same; auto ]
- | progress simpl in *
- | constructor
- | solve [ eauto ]
- | progress inversion_wf_constr
- | match goal with
- | [ H : invert_PairsConst _ _ = ?x, H' : invert_PairsConst _ _ = ?y |- _ ]
- => assert (x = y)
- by (rewrite <- H, <- H'; eapply wff_invert_PairsConst; [ eauto | eassumption ]);
- inversion_option;
- progress (subst || congruence)
- end
- | break_innermost_match_hyps_step
- | break_innermost_match_step ].
- Lemma wff_rewrite_for_const_and_op G s d opc args1 args2
- (Hwf : wff G (var1:=var1) (var2:=var2) args1 args2)
- : wff G
- (rewrite_for_const_and_op interp_op1 invert_Const1 Const1 s d opc args1)
- (rewrite_for_const_and_op interp_op2 invert_Const2 Const2 s d opc args2).
- Proof using HConst Hinterp_op Hinvert_Const. cbv [rewrite_for_const_and_op]; t_fin. Qed.
-
- Lemma wff_inline_const_and_op_genf {t} e1 e2 G
- (wf : wff G e1 e2)
- : @wff var1 var2 G t
- (inline_const_and_op_genf interp_op1 invert_Const1 Const1 e1)
- (inline_const_and_op_genf interp_op2 invert_Const2 Const2 e2).
- Proof using HConst Hinvert_Const Hinterp_op.
- unfold inline_const_and_op_genf;
- eauto using wff_rewrite_opf, wff_rewrite_for_const_and_op.
- Qed.
-
- Lemma wff_inline_const_and_op_gen {t} e1 e2
- (Hwf : wf e1 e2)
- : @wf var1 var2 t
- (inline_const_and_op_gen interp_op1 invert_Const1 Const1 e1)
- (inline_const_and_op_gen interp_op2 invert_Const2 Const2 e2).
- Proof using HConst Hinvert_Const Hinterp_op.
- unfold inline_const_and_op_gen;
- eauto using wf_rewrite_op, wff_rewrite_for_const_and_op.
- Qed.
- End with_var.
-
- Section const_unit.
- Context {interp_op1 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {interp_op2 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {var1 var2 : base_type_code -> Type}
- (OpConst1 OpConst2 : forall t, interp_base_type t -> op Unit (Tbase t))
- (HOpConst : forall t v, OpConst1 t v = OpConst2 t v)
- (Hinterp_op : forall s d opv args, interp_op1 s d opv args = interp_op2 s d opv args).
-
- Lemma wff_invert_ConstUnit
- s d opv e1 e2
- : @invert_ConstUnit _ _ _ interp_op1 var1 s d opv e1
- = @invert_ConstUnit _ _ _ interp_op2 var2 s d opv e2.
- Proof using Hinterp_op.
- cbv [invert_ConstUnit invert_ConstUnit']; destruct s; f_equal; auto.
- Qed.
-
- Lemma wff_Const {t} v G
- : wff G (@Const _ _ _ var1 OpConst1 t v) (@Const _ _ _ var2 OpConst2 t v).
- Proof using HOpConst.
- cbv [Const]; rewrite HOpConst; auto.
- Qed.
-
- Lemma wff_inline_const_and_opf {t} e1 e2 G
- (wf : wff G e1 e2)
- : @wff var1 var2 G t
- (inline_const_and_opf interp_op1 OpConst1 e1)
- (inline_const_and_opf interp_op2 OpConst2 e2).
- Proof using HOpConst Hinterp_op.
- unfold inline_const_and_opf;
- eauto using wff_inline_const_and_op_genf, wff_invert_ConstUnit, wff_Const.
- Qed.
-
- Lemma wff_inline_const_and_op {t} e1 e2
- (Hwf : wf e1 e2)
- : @wf var1 var2 t (inline_const_and_op interp_op1 OpConst1 e1) (inline_const_and_op interp_op2 OpConst2 e2).
- Proof using HOpConst Hinterp_op.
- unfold inline_const_and_op;
- eauto using wff_inline_const_and_op_gen, wff_invert_ConstUnit, wff_Const.
- Qed.
- End const_unit.
-
- Lemma Wf_InlineConstAndOpGen
- {interp_op}
- (invert_Const : forall var s d, op s d -> @exprf var s -> option (interp_flat_type interp_base_type d))
- (Const : forall var t, interp_base_type t -> @exprf var (Tbase t))
- (Hinvert_Const
- : forall var1 var2 s d opv G e1 e2,
- wff G e1 e2
- -> invert_Const var1 s d opv e1 = invert_Const var2 s d opv e2)
- (HConst
- : forall var1 var2 t v G,
- wff G (Const var1 t v) (Const var2 t v))
- {t} (e : Expr t)
- (Hwf : Wf e)
- : Wf (InlineConstAndOpGen interp_op invert_Const Const e).
- Proof using Type.
- unfold InlineConstAndOpGen;
- eauto using Wf_RewriteOp, wff_rewrite_for_const_and_op.
- Qed.
-
- Lemma Wf_InlineConstAndOp
- {interp_op}
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- {t} (e : Expr t)
- (Hwf : Wf e)
- : Wf (InlineConstAndOp interp_op OpConst e).
- Proof using Type.
- unfold InlineConstAndOp;
- eauto using Wf_RewriteOp, wff_rewrite_for_const_and_op, wff_Const.
- Qed.
- End language.
-
- Hint Resolve Wf_InlineConstAndOpGen Wf_InlineConstAndOp : wf.
-End Rewrite.
diff --git a/src/Compilers/InlineConstAndOpInterp.v b/src/Compilers/InlineConstAndOpInterp.v
deleted file mode 100644
index e464d1e8a..000000000
--- a/src/Compilers/InlineConstAndOpInterp.v
+++ /dev/null
@@ -1,161 +0,0 @@
-(** * Inline: Remove some [Let] expressions, inline constants, interpret constant operations *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.InlineConstAndOpWf.
-Require Import Crypto.Compilers.InterpProofs.
-Require Import Crypto.Compilers.Inline.
-Require Import Crypto.Compilers.InlineInterp.
-Require Import Crypto.Compilers.InlineConstAndOp.
-Require Import Crypto.Util.Sigma Crypto.Util.Prod Crypto.Util.Option.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-
-Local Open Scope ctype_scope.
-Section language.
- Context (base_type_code : Type)
- (interp_base_type : base_type_code -> Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
-
- Section with_invert.
- Context (invert_Const : forall s d, op s d -> @exprf interp_base_type s -> option (interp_flat_type d))
- (Const : forall t, interp_base_type t -> @exprf interp_base_type (Tbase t))
- (Hinvert_Const
- : forall s d opc e v, invert_Const s d opc e = Some v
- -> interp_op s d opc (interpf interp_op e) = v)
- (interpf_Const : forall t v, interpf interp_op (Const t v) = v).
-
- Lemma interpf_postprocess_for_const_and_op {t} (e : exprf t)
- : interpf interp_op
- (exprf_of_inline_directive
- (postprocess_for_const_and_op interp_op invert_Const Const e))
- = interpf interp_op e.
- Proof.
- induction e; try reflexivity; simpl in *.
- all:repeat first [ reflexivity
- | break_innermost_match_step
- | progress cbv [SmartVarVarf]
- | progress cbn [interpf exprf_of_inline_directive interpf_step LetIn.Let_In SmartVarVarf fst snd] in *
- | solve [ auto ]
- | rewrite SmartPairf_Pair
- | apply (f_equal2 (@pair _ _))
- | rewrite interpf_SmartPairf
- | rewrite SmartVarfMap_compose
- | rewrite SmartVarfMap_id
- | setoid_rewrite interpf_Const
- | erewrite ExprInversion.interpf_invert_PairsConst by eassumption ].
- Qed.
-
- Lemma interpf_inline_const_and_op_genf
- G {t} e1 e2
- (wf : @wff _ _ G t e1 e2)
- (H : forall t x x',
- List.In
- (existT (fun t : base_type_code => (exprf (Tbase t) * interp_base_type t)%type) t
- (x, x')) G
- -> interpf interp_op x = x')
- : interpf interp_op (inline_const_and_op_genf (t:=t) interp_op invert_Const Const e1)
- = interpf interp_op e2.
- Proof.
- unfold inline_const_and_op_genf;
- eapply interpf_inline_const_genf; eauto using interpf_postprocess_for_const_and_op.
- Qed.
-
- Lemma interpf_inline_const_and_op_gen
- {t} e1 e2
- (Hwf : @wf _ _ t e1 e2)
- : forall x,
- interp interp_op (inline_const_and_op_gen (t:=t) interp_op invert_Const Const e1) x
- = interp interp_op e2 x.
- Proof.
- unfold inline_const_and_op_gen;
- eapply interp_inline_const_gen; eauto using interpf_postprocess_for_const_and_op.
- Qed.
- End with_invert.
-
- Section const_unit.
- Context (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- (interp_op_OpConst : forall t v, interp_op _ _ (OpConst t v) tt = v).
-
- Lemma interpf_invert_ConstUnit s d opc e v
- (H : invert_ConstUnit interp_op opc e = Some v)
- : interp_op s d opc (interpf interp_op e) = v.
- Proof using Type.
- destruct s; simpl in *; inversion_option; subst.
- edestruct interpf; reflexivity.
- Qed.
-
- Lemma interpf_Const t v
- : interpf interp_op (Const OpConst (t:=t) v) = v.
- Proof using interp_op_OpConst.
- apply interp_op_OpConst.
- Qed.
-
- Lemma interpf_inline_const_and_opf
- G {t} e1 e2
- (wf : @wff _ _ G t e1 e2)
- (H : forall t x x',
- List.In
- (existT (fun t : base_type_code => (exprf (Tbase t) * interp_base_type t)%type) t
- (x, x')) G
- -> interpf interp_op x = x')
- : interpf interp_op (inline_const_and_opf (t:=t) interp_op OpConst e1)
- = interpf interp_op e2.
- Proof.
- unfold inline_const_and_opf;
- eapply interpf_inline_const_genf; eauto using interpf_postprocess_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
-
- Lemma interpf_inline_const_and_op
- {t} e1 e2
- (Hwf : @wf _ _ t e1 e2)
- : forall x,
- interp interp_op (inline_const_and_op (t:=t) interp_op OpConst e1) x
- = interp interp_op e2 x.
- Proof.
- unfold inline_const_and_op;
- eapply interp_inline_const_gen; eauto using interpf_postprocess_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
- End const_unit.
-
- Lemma InterpInlineConstAndOpGen
- (invert_Const : forall var s d, op s d -> @exprf var s -> option (interp_flat_type d))
- (Const : forall var t, interp_base_type t -> @exprf var (Tbase t))
-
- {t} (e : Expr t)
- (wf : Wf e)
- (Hinvert_Const
- : forall s d opc e v,
- invert_Const _ s d opc e = Some v
- -> interp_op s d opc (interpf interp_op e) = v)
- (interpf_Const : forall t v, interpf interp_op (Const _ t v) = v)
- : forall x, Interp interp_op (InlineConstAndOpGen interp_op invert_Const Const e) x = Interp interp_op e x.
- Proof using Type.
- eapply InterpInlineConstGen;
- eauto using interpf_postprocess_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
-
- Lemma InterpInlineConstAndOp
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- {t} (e : Expr t)
- (wf : Wf e)
- (interp_op_OpConst : forall t v, interp_op _ _ (OpConst t v) tt = v)
- : forall x, Interp interp_op (InlineConstAndOp interp_op OpConst e) x = Interp interp_op e x.
- Proof using Type.
- eapply InterpInlineConstGen;
- eauto using interpf_postprocess_for_const_and_op, interpf_invert_ConstUnit, interpf_Const.
- Qed.
-End language.
-
-(*Hint Rewrite @InterpInlineConst @interp_inline_const @interpf_inline_constf using solve_wf_side_condition : reflective_interp.*)
diff --git a/src/Compilers/InlineConstAndOpWf.v b/src/Compilers/InlineConstAndOpWf.v
deleted file mode 100644
index 0de978e90..000000000
--- a/src/Compilers/InlineConstAndOpWf.v
+++ /dev/null
@@ -1,219 +0,0 @@
-(** * Inline: Remove some [Let] expressions, inline constants, interpret constant operations *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.Inline.
-Require Import Crypto.Compilers.InlineWf.
-Require Import Crypto.Compilers.InlineConstAndOp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sigma Crypto.Util.Prod.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}.
-
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation Tbase := (@Tbase base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
-
- Local Hint Constructors Wf.wff.
-
- Local Notation wff_inline_directive G x y :=
- (wff G (exprf_of_inline_directive x) (exprf_of_inline_directive y)
- /\ (fun x' y'
- => match x', y' with
- | default_inline _ _, default_inline _ _
- | no_inline _ _, no_inline _ _
- | inline _ _, inline _ _
- | partial_inline _ _ _ _, partial_inline _ _ _ _
- => True
- | default_inline _ _, _
- | no_inline _ _, _
- | inline _ _, _
- | partial_inline _ _ _ _, _
- => False
- end) x y).
-
-
- Section with_var.
- Context {interp_op1 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {interp_op2 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {var1 var2 : base_type_code -> Type}
- {invert_Const1 : forall s d, op s d -> @exprf var1 s -> option (interp_flat_type interp_base_type d)}
- {invert_Const2 : forall s d, op s d -> @exprf var2 s -> option (interp_flat_type interp_base_type d)}
- {Const1 : forall t, interp_base_type t -> @exprf var1 (Tbase t)}
- {Const2 : forall t, interp_base_type t -> @exprf var2 (Tbase t)}
- (Hinterp_op : forall s d opv args, interp_op1 s d opv args = interp_op2 s d opv args)
- (Hinvert_Const : forall s d opv G e1 e2, wff G e1 e2 -> invert_Const1 s d opv e1 = invert_Const2 s d opv e2)
- (HConst : forall t v G, wff G (Const1 t v) (Const2 t v)).
-
- Local Ltac t_fin :=
- repeat first [ intro
- | progress inversion_sigma
- | progress inversion_prod
- | exfalso; assumption
- | exact I
- | progress destruct_head'_prod
- | progress destruct_head'_and
- | progress destruct_head'_sig
- | progress subst
- | rewrite Hinterp_op
- | rewrite SmartPairfSmartVarVarf_SmartVarf
- | rewrite SmartPairf_Pair
- | tauto
- | solve [ auto with nocore
- | eapply (@wff_SmartVarVarf _ _ _ _ _ _ (_ * _)); auto
- | eapply wff_SmartVarVarf; eauto with nocore
- | eapply wff_SmartPairf; eauto with nocore
- | eapply wff_app_post, wff_SmartVarf
- | apply wff_SmartPairf_SmartVarfMap_same; auto ]
- | progress simpl in *
- | constructor
- | solve [ eauto
- | eauto using wff_app_pre, wff_app_post
- | eapply wff_in_impl_Proper;
- [ solve [ eauto using wff_SmartVarf ]
- | intros *;
- repeat setoid_rewrite List.in_app_iff;
- tauto ] ]
- | progress inversion_wf_constr
- | match goal with
- | [ H : invert_PairsConst _ _ = ?x, H' : invert_PairsConst _ _ = ?y |- _ ]
- => assert (x = y)
- by (rewrite <- H, <- H'; eapply wff_invert_PairsConst; [ eauto | eassumption ]);
- inversion_option;
- progress (subst || congruence)
- end
- | break_innermost_match_hyps_step
- | break_innermost_match_step ].
-
- Lemma wff_postprocess_for_const_and_op
- G t e1 e2
- (Hwf : wff (t:=t) G e1 e2)
- : wff_inline_directive
- G
- (postprocess_for_const_and_op interp_op1 invert_Const1 Const1 e1)
- (postprocess_for_const_and_op interp_op2 invert_Const2 Const2 e2).
- Proof using Hinvert_Const Hinterp_op HConst.
- induction Hwf; split; cbv beta;
- repeat match goal with
- | [ H : wff _ _ _ /\ _ |- match _ with _ => _ end _ ] => destruct H as [_ ?]
- end.
- { t_fin. }
- { t_fin. }
- { t_fin. }
- { t_fin. }
- { destruct IHHwf as [? _]. (* for speed *)
- t_fin. }
- { t_fin. }
- { t_fin. }
- { t_fin. }
- { t_fin. }
- { t_fin. }
- Qed.
-
- Lemma wff_inline_const_and_op_genf {t} e1 e2 G G'
- (H : forall t x x', List.In (existT (fun t : base_type_code => (exprf (Tbase t) * exprf (Tbase t))%type) t (x, x')) G'
- -> wff G x x')
- (wf : wff G' e1 e2)
- : @wff var1 var2 G t (inline_const_and_op_genf interp_op1 invert_Const1 Const1 e1) (inline_const_and_op_genf interp_op2 invert_Const2 Const2 e2).
- Proof using HConst Hinvert_Const Hinterp_op.
- unfold inline_const_and_op_genf;
- eauto using wff_inline_const_genf, wff_postprocess_for_const_and_op.
- Qed.
-
- Lemma wff_inline_const_and_op_gen {t} e1 e2
- (Hwf : wf e1 e2)
- : @wf var1 var2 t (inline_const_and_op_gen interp_op1 invert_Const1 Const1 e1) (inline_const_and_op_gen interp_op2 invert_Const2 Const2 e2).
- Proof using HConst Hinvert_Const Hinterp_op.
- unfold inline_const_and_op_gen;
- eauto using wf_inline_const_gen, wff_postprocess_for_const_and_op.
- Qed.
- End with_var.
-
- Section const_unit.
- Context {interp_op1 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {interp_op2 : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {var1 var2 : base_type_code -> Type}
- (OpConst1 OpConst2 : forall t, interp_base_type t -> op Unit (Tbase t))
- (HOpConst : forall t v, OpConst1 t v = OpConst2 t v)
- (Hinterp_op : forall s d opv args, interp_op1 s d opv args = interp_op2 s d opv args).
-
- Lemma wff_invert_ConstUnit
- s d opv e1 e2
- : @invert_ConstUnit _ _ _ interp_op1 var1 s d opv e1
- = @invert_ConstUnit _ _ _ interp_op2 var2 s d opv e2.
- Proof using Hinterp_op.
- cbv [invert_ConstUnit invert_ConstUnit']; destruct s; f_equal; auto.
- Qed.
-
- Lemma wff_Const {t} v G
- : wff G (@Const _ _ _ var1 OpConst1 t v) (@Const _ _ _ var2 OpConst2 t v).
- Proof using HOpConst.
- cbv [Const]; rewrite HOpConst; auto.
- Qed.
-
- Lemma wff_inline_const_and_opf {t} e1 e2 G G'
- (H : forall t x x', List.In (existT (fun t : base_type_code => (exprf (Tbase t) * exprf (Tbase t))%type) t (x, x')) G'
- -> wff G x x')
- (wf : wff G' e1 e2)
- : @wff var1 var2 G t (inline_const_and_opf interp_op1 OpConst1 e1) (inline_const_and_opf interp_op2 OpConst2 e2).
- Proof using HOpConst Hinterp_op.
- unfold inline_const_and_opf;
- eauto using wff_inline_const_and_op_genf, wff_invert_ConstUnit, wff_Const.
- Qed.
-
- Lemma wff_inline_const_and_op {t} e1 e2
- (Hwf : wf e1 e2)
- : @wf var1 var2 t (inline_const_and_op interp_op1 OpConst1 e1) (inline_const_and_op interp_op2 OpConst2 e2).
- Proof using HOpConst Hinterp_op.
- unfold inline_const_and_op;
- eauto using wff_inline_const_and_op_gen, wff_invert_ConstUnit, wff_Const.
- Qed.
- End const_unit.
-
- Lemma Wf_InlineConstAndOpGen
- {interp_op}
- (invert_Const : forall var s d, op s d -> @exprf var s -> option (interp_flat_type interp_base_type d))
- (Const : forall var t, interp_base_type t -> @exprf var (Tbase t))
- (Hinvert_Const
- : forall var1 var2 s d opv G e1 e2,
- wff G e1 e2
- -> invert_Const var1 s d opv e1 = invert_Const var2 s d opv e2)
- (HConst
- : forall var1 var2 t v G,
- wff G (Const var1 t v) (Const var2 t v))
- {t} (e : Expr t)
- (Hwf : Wf e)
- : Wf (InlineConstAndOpGen interp_op invert_Const Const e).
- Proof using Type.
- unfold InlineConstAndOpGen;
- eauto using Wf_InlineConstGen, wff_postprocess_for_const_and_op.
- Qed.
-
- Lemma Wf_InlineConstAndOp
- {interp_op}
- (OpConst : forall t, interp_base_type t -> op Unit (Tbase t))
- {t} (e : Expr t)
- (Hwf : Wf e)
- : Wf (InlineConstAndOp interp_op OpConst e).
- Proof using Type.
- unfold InlineConstAndOp;
- eauto using Wf_InlineConstGen, wff_postprocess_for_const_and_op, wff_Const.
- Qed.
-End language.
-
-Hint Resolve Wf_InlineConstAndOpGen Wf_InlineConstAndOp : wf.
diff --git a/src/Compilers/InlineInterp.v b/src/Compilers/InlineInterp.v
deleted file mode 100644
index caef4245a..000000000
--- a/src/Compilers/InlineInterp.v
+++ /dev/null
@@ -1,137 +0,0 @@
-(** * Inline: Remove some [Let] expressions *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.InlineWf.
-Require Import Crypto.Compilers.InterpProofs.
-Require Import Crypto.Compilers.Inline.
-Require Import Crypto.Util.Sigma Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-
-
-Local Open Scope ctype_scope.
-Section language.
- Context (base_type_code : Type).
- Context (interp_base_type : base_type_code -> Type).
- Context (op : flat_type base_type_code -> flat_type base_type_code -> Type).
- Context (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
-
- Local Hint Extern 1 => eapply interpf_SmartVarVarf.
-
- Local Ltac t_fin_step :=
- match goal with
- | _ => reflexivity
- | _ => progress simpl in *
- | _ => progress unfold postprocess_for_const in *
- | _ => progress intros
- | _ => progress inversion_sigma
- | _ => progress inversion_prod
- | _ => solve [ intuition eauto ]
- | _ => apply (f_equal (interp_op _ _ _))
- | _ => apply (f_equal2 (@pair _ _))
- | _ => progress specialize_by assumption
- | _ => progress subst
- | [ H : context[List.In _ (_ ++ _)] |- _ ] => setoid_rewrite List.in_app_iff in H
- | [ H : _ = _ :> inline_directive _ |- _ ]
- => apply (f_equal exprf_of_inline_directive) in H
- | [ H : exprf_of_inline_directive _ = _ |- _ ]
- => apply (f_equal (interpf interp_op)) in H
- | [ H : @fst ?A ?B ?x = _, H' : context H'T[@fst ?A' ?B' ?x] |- _ ]
- => let H'T' := context H'T[@fst A B x] in
- progress change H'T' in H'
- | [ H : @snd ?A ?B ?x = _, H' : context H'T[@snd ?A' ?B' ?x] |- _ ]
- => let H'T' := context H'T[@snd A B x] in
- progress change H'T' in H'
- | [ H : or _ _ |- _ ] => destruct H
- | _ => progress break_match
- | _ => rewrite <- !surjective_pairing
- | [ H : ?x = _, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : _ |- _ ] => rewrite H; []
- | [ H : _, H' : _ |- _ ] => rewrite H in H' by fail
- | [ H : _ |- _ ] => apply H; solve [ repeat t_fin_step ]
- | [ H : _ |- _ ] => rewrite H; solve [ repeat t_fin_step ]
- | _ => solve [ eapply WfProofs.flatten_binding_list_interpf_SmartPairf_same; eauto ]
- end.
- Local Ltac t_fin := repeat t_fin_step.
-
- Lemma interpf_inline_const_genf postprocess G {t} e1 e2
- (wf : @wff _ _ G t e1 e2)
- (Hpostprocess : forall t e, interpf interp_op (exprf_of_inline_directive (postprocess t e)) = interpf interp_op e)
- (H : forall t x x',
- List.In
- (existT (fun t : base_type_code => (exprf (Tbase t) * interp_base_type t)%type) t
- (x, x')) G
- -> interpf interp_op x = x')
- : interpf interp_op (inline_const_genf postprocess e1) = interpf interp_op e2.
- Proof using Type.
- clear -wf H Hpostprocess.
- induction wf; t_fin.
- Qed.
-
- Lemma interpf_postprocess_for_const is_const t e
- : interpf interp_op (exprf_of_inline_directive (postprocess_for_const is_const t e)) = interpf interp_op e.
- Proof using Type.
- unfold postprocess_for_const; t_fin.
- Qed.
-
- Local Hint Resolve interpf_postprocess_for_const.
-
- Lemma interpf_inline_constf is_const G {t} e1 e2
- (wf : @wff _ _ G t e1 e2)
- (H : forall t x x',
- List.In
- (existT (fun t : base_type_code => (exprf (Tbase t) * interp_base_type t)%type) t
- (x, x')) G
- -> interpf interp_op x = x')
- : interpf interp_op (inline_constf is_const e1) = interpf interp_op e2.
- Proof using Type. eapply interpf_inline_const_genf; eauto. Qed.
-
- Local Hint Resolve interpf_inline_constf.
-
- Lemma interp_inline_const_gen postprocess {t} e1 e2
- (wf : @wf _ _ t e1 e2)
- (Hpostprocess : forall t e, interpf interp_op (exprf_of_inline_directive (postprocess t e)) = interpf interp_op e)
- : forall x, interp interp_op (inline_const_gen postprocess e1) x = interp interp_op e2 x.
- Proof using Type.
- destruct wf.
- simpl in *; intro; eapply (interpf_inline_const_genf postprocess); eauto.
- Qed.
-
- Local Hint Resolve interp_inline_const_gen.
-
- Lemma interp_inline_const is_const {t} e1 e2
- (wf : @wf _ _ t e1 e2)
- : forall x, interp interp_op (inline_const is_const e1) x = interp interp_op e2 x.
- Proof using Type.
- eapply interp_inline_const_gen; eauto.
- Qed.
-
- Lemma InterpInlineConstGen postprocess {t} (e : Expr t)
- (wf : Wf e)
- (Hpostprocess : forall t e, interpf interp_op (exprf_of_inline_directive (postprocess _ t e)) = interpf interp_op e)
- : forall x, Interp interp_op (InlineConstGen postprocess e) x = Interp interp_op e x.
- Proof using Type.
- unfold Interp, InlineConst.
- eapply (interp_inline_const_gen (postprocess _)); simpl; intuition.
- Qed.
-
- Lemma InterpInlineConst is_const {t} (e : Expr t)
- (wf : Wf e)
- : forall x, Interp interp_op (InlineConst is_const e) x = Interp interp_op e x.
- Proof using Type.
- eapply InterpInlineConstGen; eauto.
- Qed.
-End language.
-
-Hint Rewrite @InterpInlineConst @interp_inline_const @interpf_inline_constf using solve_wf_side_condition : reflective_interp.
diff --git a/src/Compilers/InlineWf.v b/src/Compilers/InlineWf.v
deleted file mode 100644
index a9fc959d0..000000000
--- a/src/Compilers/InlineWf.v
+++ /dev/null
@@ -1,231 +0,0 @@
-(** * Inline: Remove some [Let] expressions *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.TypeInversion.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Compilers.Inline.
-Require Import Crypto.Util.Tactics.SpecializeBy Crypto.Util.Tactics.DestructHead Crypto.Util.Sigma Crypto.Util.Prod.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.UniquePose.
-Require Import Crypto.Util.Tactics.SplitInContext.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation Tbase := (@Tbase base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
-
- Local Notation wff_inline_directive G x y :=
- (wff G (exprf_of_inline_directive x) (exprf_of_inline_directive y)
- /\ (fun x' y'
- => match x', y' with
- | default_inline _ _, default_inline _ _
- | no_inline _ _, no_inline _ _
- | inline _ _, inline _ _
- | partial_inline _ _ _ _, partial_inline _ _ _ _
- => True
- | default_inline _ _, _
- | no_inline _ _, _
- | inline _ _, _
- | partial_inline _ _ _ _, _
- => False
- end) x y).
-
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}.
-
- Local Hint Constructors Wf.wff.
- Local Hint Resolve List.in_app_or List.in_or_app.
-
- Local Hint Constructors or.
- Local Hint Constructors Wf.wff.
- Local Hint Extern 1 => progress unfold List.In in *.
- Local Hint Resolve wff_in_impl_Proper.
- Local Hint Resolve wff_SmartVarf.
- Local Hint Resolve wff_SmartVarVarf.
-
- Local Ltac t_fin :=
- repeat first [ intro
- | progress inversion_sigma
- | progress inversion_prod
- | tauto
- | progress subst
- | solve [ auto with nocore
- | eapply (@wff_SmartVarVarf _ _ _ _ _ _ (_ * _)); auto
- | eapply wff_SmartVarVarf; eauto with nocore
- | eapply wff_SmartPairf; eauto with nocore ]
- | progress simpl in *
- | constructor
- | solve [ eauto ] ].
-
- Local Ltac invert_inline_directive' i :=
- preinvert_one_type i;
- intros ? i;
- destruct i;
- try exact I.
- Local Ltac invert_inline_directive :=
- match goal with
- | [ i : inline_directive _ |- _ ] => invert_inline_directive' i
- end.
-
- (** XXX TODO: Clean up this proof *)
- Lemma wff_inline_const_genf postprocess1 postprocess2 {t} e1 e2 G G'
- (H : forall t x x', List.In (existT (fun t : base_type_code => (exprf (Tbase t) * exprf (Tbase t))%type) t (x, x')) G'
- -> wff G x x')
- (wf : wff G' e1 e2)
- (wf_postprocess : forall G t e1 e2,
- wff G e1 e2
- -> wff_inline_directive G (postprocess1 t e1) (postprocess2 t e2))
- : @wff var1 var2 G t (inline_const_genf postprocess1 e1) (inline_const_genf postprocess2 e2).
- Proof using Type.
- revert dependent G; induction wf; simpl in *; auto; intros; [].
- repeat match goal with
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => setoid_rewrite List.in_app_iff in H
- | [ |- context[postprocess1 ?t ?e1] ]
- => match goal with
- | [ |- context[postprocess2 t ?e2] ]
- => specialize (fun G => wf_postprocess G t e1 e2);
- generalize dependent (postprocess1 t e1);
- generalize dependent (postprocess2 t e2)
- end
- | _ => intro
- end.
- repeat match goal with
- | [ H : forall G : list _, _ |- wff ?G' _ _ ]
- => unique pose proof (H G')
- | [ H : forall x y (G : list _), _ |- wff ?G' _ _ ]
- => unique pose proof (fun x y => H x y G')
- | [ H : forall x1 x2, (forall t x x', _ \/ List.In _ ?G -> wff ?G0 x x') -> _,
- H' : forall t1 x1 x1', List.In _ ?G -> wff ?G0 x1 x1' |- _ ]
- => unique pose proof (fun x1 x2 f => H x1 x2 (fun t x x' pf => match pf with or_introl pf => f t x x' pf | or_intror pf => H' t x x' pf end))
- end.
- repeat match goal with
- | _ => exact I
- | [ H : forall x1 : unit, _ |- _ ] => specialize (H tt)
- | [ H : False |- _ ] => exfalso; assumption
- | _ => progress subst
- | _ => progress inversion_sigma
- | _ => progress inversion_prod
- | _ => progress destruct_head' and
- | _ => inversion_wf_step; destruct_head' sig; progress subst
- | _ => progress specialize_by_assumption
- | _ => progress break_match
- | _ => progress invert_inline_directive
- | [ |- context[match ?e with _ => _ end] ]
- => invert_one_expr e
- | _ => progress destruct_head' or
- | _ => progress simpl in *
- | _ => intro
- | _ => progress split_and
- | [ H : wff _ TT _ |- _ ] => solve [ inversion H ]
- | [ H : wff _ (Var _ _) _ |- _ ] => solve [ inversion H ]
- | [ H : wff _ (Op _ _) _ |- _ ] => solve [ inversion H ]
- | [ H : wff _ (LetIn _ _) _ |- _ ] => solve [ inversion H ]
- | [ H : wff _ (Pair _ _) _ |- _ ] => solve [ inversion H ]
- end;
- repeat first [ progress specialize_by tauto
- | progress specialize_by auto
- | solve [ auto ] ];
- try (constructor; auto; intros).
- { match goal with H : _ |- _ => apply H end.
- intros; destruct_head or; t_fin. }
- { match goal with H : _ |- _ => apply H end.
- intros; destruct_head or; t_fin. }
- { match goal with H : _ |- _ => apply H end.
- intros; destruct_head or; t_fin. }
- { match goal with H : _ |- _ => apply H end.
- intros; destruct_head' or; t_fin. }
- { match goal with H : _ |- _ => apply H end.
- intros; destruct_head or; t_fin. }
- { match goal with H : _ |- _ => apply H end.
- intros; destruct_head or; t_fin. }
- { match goal with H : _ |- _ => apply H end.
- intros; destruct_head or; t_fin. }
- Qed.
-
- Lemma wff_postprocess_for_const is_const G t
- (e1 : @exprf var1 t)
- (e2 : @exprf var2 t)
- (Hwf : wff G e1 e2)
- : wff_inline_directive G (postprocess_for_const is_const t e1) (postprocess_for_const is_const t e2).
- Proof using Type.
- destruct e1; unfold postprocess_for_const;
- repeat first [ progress subst
- | intro
- | progress destruct_head' sig
- | progress destruct_head' and
- | progress inversion_sigma
- | progress inversion_option
- | progress inversion_prod
- | progress destruct_head' False
- | progress simpl in *
- | progress unfold SmartMap.SmartPairf
- | progress invert_expr
- | progress inversion_wf
- | progress break_innermost_match_step
- | discriminate
- | congruence
- | solve [ auto ] ].
- Qed.
-
- Local Hint Resolve wff_postprocess_for_const.
-
- Lemma wff_inline_constf is_const {t} e1 e2 G G'
- (H : forall t x x', List.In (existT (fun t : base_type_code => (exprf (Tbase t) * exprf (Tbase t))%type) t (x, x')) G'
- -> wff G x x')
- (wf : wff G' e1 e2)
- : @wff var1 var2 G t (inline_constf is_const e1) (inline_constf is_const e2).
- Proof using Type. eapply wff_inline_const_genf; eauto. Qed.
-
- Lemma wf_inline_const_gen postprocess1 postprocess2 {t} e1 e2
- (Hwf : wf e1 e2)
- (wf_postprocess : forall G t e1 e2,
- wff G e1 e2
- -> wff_inline_directive G (postprocess1 t e1) (postprocess2 t e2))
- : @wf var1 var2 t (inline_const_gen postprocess1 e1) (inline_const_gen postprocess2 e2).
- Proof using Type.
- destruct Hwf; constructor; intros.
- eapply wff_inline_const_genf; eauto using wff_SmartVarVarf_nil.
- Qed.
-
- Lemma wf_inline_const is_const {t} e1 e2
- (Hwf : wf e1 e2)
- : @wf var1 var2 t (inline_const is_const e1) (inline_const is_const e2).
- Proof using Type. eapply wf_inline_const_gen; eauto. Qed.
- End with_var.
-
- Lemma Wf_InlineConstGen postprocess {t} (e : Expr t)
- (Hwf : Wf e)
- (Hpostprocess : forall var1 var2 G t e1 e2,
- wff G e1 e2
- -> wff_inline_directive G (postprocess var1 t e1) (postprocess var2 t e2))
- : Wf (InlineConstGen postprocess e).
- Proof using Type.
- intros var1 var2.
- apply (@wf_inline_const_gen var1 var2 (postprocess _) (postprocess _) t (e _) (e _)); simpl; auto.
- Qed.
-
- Lemma Wf_InlineConst is_const {t} (e : Expr t)
- (Hwf : Wf e)
- : Wf (InlineConst is_const e).
- Proof using Type.
- intros var1 var2.
- apply (@wf_inline_const var1 var2 is_const t (e _) (e _)); simpl.
- apply Hwf.
- Qed.
-End language.
-
-Hint Resolve Wf_InlineConstGen Wf_InlineConst : wf.
diff --git a/src/Compilers/InputSyntax.v b/src/Compilers/InputSyntax.v
deleted file mode 100644
index af2a6f277..000000000
--- a/src/Compilers/InputSyntax.v
+++ /dev/null
@@ -1,252 +0,0 @@
-(** * PHOAS Representation of Gallina which allows exact denotation *)
-Require Import Coq.Strings.String.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.InterpProofs.
-Require Import Crypto.Util.Tuple.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Notations.
-
-(** We parameterize the language over a type of basic type codes (for
- things like [Z], [word], [bool]), as well as a type of n-ary
- operations returning one value, and n-ary operations returning two
- values. *)
-Local Open Scope ctype_scope.
-Section language.
- Context (base_type_code : Type).
-
- Local Notation flat_type := (flat_type base_type_code).
- Inductive type := Tflat (A : flat_type) | Arrow (A : flat_type) (B : type).
-
- Section expr_param.
- Context (interp_base_type : base_type_code -> Type).
- Context (op : flat_type (* input tuple *) -> flat_type (* output type *) -> Type).
- Local Notation interp_flat_type_gen := interp_flat_type.
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
-
- Fixpoint interp_type (t : type) :=
- match t with
- | Tflat A => interp_flat_type A
- | Arrow A B => (interp_flat_type A -> interp_type B)%type
- end.
-
- Section expr.
- Context {var : flat_type -> Type}.
-
- (** N.B. [Let] destructures pairs *)
- Inductive exprf : flat_type -> Type :=
- | Const {t : flat_type} : interp_flat_type t -> exprf t
- | Var {t} : var t -> exprf t
- | Op {t1 tR} : op t1 tR -> exprf t1 -> exprf tR
- | LetIn : forall {tx}, exprf tx -> forall {tC}, (var tx -> exprf tC) -> exprf tC
- | Pair : forall {t1}, exprf t1 -> forall {t2}, exprf t2 -> exprf (Prod t1 t2)
- | MatchPair : forall {t1 t2}, exprf (Prod t1 t2) -> forall {tC}, (var t1 -> var t2 -> exprf tC) -> exprf tC.
- Inductive expr : type -> Type :=
- | Return {T} : exprf T -> expr (Tflat T)
- | Abs {src dst} : (var src -> expr dst) -> expr (Arrow src dst).
-
- Definition Fst {t1 t2} (v : exprf (Prod t1 t2)) : exprf t1 := MatchPair v (fun x y => Var x).
- Definition Snd {t1 t2} (v : exprf (Prod t1 t2)) : exprf t2 := MatchPair v (fun x y => Var y).
- End expr.
-
- Definition Expr (t : type) := forall var, @expr var t.
-
- Section interp.
- Context (interp_op : forall src dst, op src dst -> interp_flat_type src -> interp_flat_type dst).
-
- Fixpoint interpf {t} (e : @exprf interp_flat_type t) : interp_flat_type t
- := match e in exprf t return interp_flat_type t with
- | Const _ x => x
- | Var _ x => x
- | Op _ _ op args => @interp_op _ _ op (@interpf _ args)
- | LetIn _ ex _ eC => dlet x := @interpf _ ex in @interpf _ (eC x)
- | Pair _ ex _ ey => (@interpf _ ex, @interpf _ ey)
- | MatchPair _ _ ex _ eC => match @interpf _ ex with pair x y => @interpf _ (eC x y) end
- end.
- Fixpoint interp {t} (e : @expr interp_flat_type t) : interp_type t
- := match e in expr t return interp_type t with
- | Return _ v => interpf v
- | Abs _ _ f => fun x => @interp _ (f x)
- end.
-
- Definition Interp {t} (E : Expr t) : interp_type t := interp (E _).
- End interp.
-
- Section compile.
- Context {var : base_type_code -> Type}
- (make_const : forall t, interp_base_type t -> op Unit (Tbase t)).
-
- Fixpoint compilet (t : type) : Syntax.type base_type_code
- := Syntax.Arrow
- match t with
- | Tflat T => Unit
- | Arrow A (Tflat B) => A
- | Arrow A B
- => A * domain (compilet B)
- end%ctype
- match t with
- | Tflat T => T
- | Arrow A B => codomain (compilet B)
- end.
-
- Fixpoint SmartConst (t : flat_type) : interp_flat_type t -> Syntax.exprf base_type_code op (var:=var) t
- := match t return interp_flat_type t -> Syntax.exprf _ _ t with
- | Unit => fun _ => TT
- | Tbase _ => fun v => Syntax.Op (make_const _ v) TT
- | Prod _ _ => fun v => Syntax.Pair (@SmartConst _ (fst v))
- (@SmartConst _ (snd v))
- end.
-
- Fixpoint compilef {t} (e : @exprf (interp_flat_type_gen var) t) : @Syntax.exprf base_type_code op var t
- := match e in exprf t return @Syntax.exprf _ _ _ t with
- | Const _ x => @SmartConst _ x
- | Var _ x => SmartMap.SmartVarf x
- | Op _ _ op args => Syntax.Op op (@compilef _ args)
- | LetIn _ ex _ eC => Syntax.LetIn (@compilef _ ex) (fun x => @compilef _ (eC x))
- | Pair _ ex _ ey => Syntax.Pair (@compilef _ ex) (@compilef _ ey)
- | MatchPair _ _ ex _ eC => Syntax.LetIn (@compilef _ ex) (fun xy => @compilef _ (eC (fst xy) (snd xy)))
- end.
-
- (* ugh, so much manual annotation *)
- Fixpoint compile {t} (e : @expr (interp_flat_type_gen var) t) : @Syntax.expr base_type_code op var (compilet t)
- := match e in expr t return @Syntax.expr _ _ _ (compilet t) with
- | Return _ v => Syntax.Abs (fun _ => compilef v)
- | Abs src dst f
- => let res := fun x => @compile _ (f x) in
- match dst
- return (_ -> Syntax.expr _ _ (compilet dst))
- -> Syntax.expr _ _ (compilet (Arrow src dst))
- with
- | Tflat T
- => fun resf => Syntax.Abs (fun x => invert_Abs (resf x) tt)
- | Arrow A B as dst'
- => match compilet dst' as cdst
- return (_ -> Syntax.expr _ _ cdst)
- -> Syntax.expr _ _ (Syntax.Arrow
- (_ * domain cdst)
- (codomain cdst))
- with
- | Syntax.Arrow A' B'
- => fun resf => Syntax.Abs (fun x : interp_flat_type_gen var (_ * _)
- => invert_Abs (resf (fst x)) (snd x))
- end
- end res
- end.
- End compile.
-
- Definition Compile
- (make_const : forall t, interp_base_type t -> op Unit (Tbase t))
- {t} (e : Expr t) : Syntax.Expr base_type_code op (compilet t)
- := fun var => compile make_const (e _).
-
- Section compile_correct.
- Context (make_const : forall t, interp_base_type t -> op Unit (Tbase t))
- (interp_op : forall src dst, op src dst -> interp_flat_type src -> interp_flat_type dst)
- (make_const_correct : forall T v, interp_op Unit (Tbase T) (make_const T v) tt = v).
-
- Lemma SmartConst_correct t v
- : Syntax.interpf interp_op (SmartConst make_const t v) = v.
- Proof using Type*.
- induction t; try destruct v; simpl in *; congruence.
- Qed.
-
- Lemma compilef_correct {t} (e : @exprf interp_flat_type t)
- : Syntax.interpf interp_op (compilef make_const e) = interpf interp_op e.
- Proof using Type*.
- induction e;
- repeat match goal with
- | _ => reflexivity
- | _ => progress unfold LetIn.Let_In
- | _ => progress simpl in *
- | _ => rewrite interpf_SmartVarf
- | _ => rewrite SmartConst_correct
- | _ => rewrite <- surjective_pairing
- | _ => progress rewrite_hyp *
- | [ |- context[let (x, y) := ?v in _] ]
- => rewrite (surjective_pairing v); cbv beta iota
- end.
- Qed.
-
- Lemma compile_flat_correct {T} (e : expr (Tflat T))
- : forall x, Syntax.interp interp_op (compile make_const e) x = interp interp_op e.
- Proof using Type*.
- intros []; simpl.
- let G := match goal with |- ?G => G end in
- let G := match (eval pattern T, e in G) with ?G _ _ => G end in
- refine match e in expr t return match t return expr t -> _ with
- | Tflat T => G T
- | _ => fun _ => True
- end e
- with
- | Return _ _ => _
- | Abs _ _ _ => I
- end; simpl.
- apply compilef_correct.
- Qed.
-
- Lemma Compile_flat_correct_flat {T} (e : Expr (Tflat T))
- : forall x, Syntax.Interp interp_op (Compile make_const e) x = Interp interp_op e.
- Proof using Type*. apply compile_flat_correct. Qed.
-
- Lemma Compile_correct {src dst} (e : @Expr (Arrow src (Tflat dst)))
- : forall x, Syntax.Interp interp_op (Compile make_const e) x = Interp interp_op e x.
- Proof using Type*.
- unfold Interp, Compile, Syntax.Interp; simpl.
- pose (e interp_flat_type) as E.
- repeat match goal with |- context[e ?f] => change (e f) with E end.
- clearbody E; clear e.
- let G := match goal with |- ?G => G end in
- let G := match (eval pattern src, dst, E in G) with ?G _ _ _ => G end in
- refine match E in expr t return match t return expr t -> _ with
- | Arrow src (Tflat dst) => G src dst
- | _ => fun _ => True
- end E
- with
- | Abs src dst e
- => match dst
- return (forall e : _ -> expr dst,
- match dst return expr (Arrow src dst) -> _ with
- | Tflat dst => G src dst
- | _ => fun _ => True
- end (Abs e))
- with
- | Tflat _
- => fun e0 x
- => _
- | Arrow _ _ => fun _ => I
- end e
- | Return _ _ => I
- end; simpl.
- refine match e0 x as e0x in expr t
- return match t return expr t -> _ with
- | Tflat _
- => fun e0x
- => Syntax.interpf _ (invert_Abs (compile _ e0x) _)
- = interp _ e0x
- | _ => fun _ => True
- end e0x
- with
- | Abs _ _ _ => I
- | Return _ _ => _
- end; simpl.
- apply compilef_correct.
- Qed.
- End compile_correct.
- End expr_param.
-End language.
-
-Global Arguments Arrow {_} _ _.
-Global Arguments Tflat {_} _.
-Global Arguments Const {_ _ _ _ _} _.
-Global Arguments Var {_ _ _ _ _} _.
-Global Arguments Op {_ _ _ _ _ _} _ _.
-Global Arguments LetIn {_ _ _ _ _} _ {_} _.
-Global Arguments MatchPair {_ _ _ _ _ _} _ {_} _.
-Global Arguments Fst {_ _ _ _ _ _} _.
-Global Arguments Snd {_ _ _ _ _ _} _.
-Global Arguments Pair {_ _ _ _ _} _ {_} _.
-Global Arguments Return {_ _ _ _ _} _.
-Global Arguments Abs {_ _ _ _ _ _} _.
-Global Arguments Compile {_ _ _} make_const {t} _ _.
diff --git a/src/Compilers/InterpByIso.v b/src/Compilers/InterpByIso.v
deleted file mode 100644
index a2263364b..000000000
--- a/src/Compilers/InterpByIso.v
+++ /dev/null
@@ -1,33 +0,0 @@
-(** * PHOAS interpretation function for any retract of [var:=interp_base_type] *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.SmartMap.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst)
- {var : base_type_code -> Type}
- (var_of_interp : forall t, interp_base_type t -> var t)
- (interp_of_var : forall t, var t -> interp_base_type t)
- (var_is_retract : forall t x, interp_of_var t (var_of_interp t x) = x).
-
- Fixpoint interpf_retr {t} (e : @exprf base_type_code op var t)
- : interp_flat_type interp_base_type t
- := match e in exprf _ _ t return interp_flat_type interp_base_type t with
- | TT => tt
- | Var t v => interp_of_var _ v
- | Op t1 tR opc args => interp_op _ _ opc (@interpf_retr _ args)
- | LetIn tx ex tC eC
- => let ev := @interpf_retr _ ex in
- @interpf_retr _ (eC (SmartVarfMap var_of_interp ev))
- | Pair tx ex ty ey => (@interpf_retr _ ex, @interpf_retr _ ey)
- end.
-
- Definition interp_retr {t} (e : @expr base_type_code op var t)
- : interp_type interp_base_type t
- := fun x => interpf_retr (invert_Abs e (SmartVarfMap var_of_interp x)).
-End language.
-
-Global Arguments interp_retr _ _ _ _ _ _ _ _ !_ / _ .
diff --git a/src/Compilers/InterpByIsoProofs.v b/src/Compilers/InterpByIsoProofs.v
deleted file mode 100644
index 98e700738..000000000
--- a/src/Compilers/InterpByIsoProofs.v
+++ /dev/null
@@ -1,117 +0,0 @@
-(** * PHOAS interpretation function for any retract of [var:=interp_base_type] *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InterpByIso.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst).
-
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation interpf_retr := (@interpf_retr base_type_code op interp_base_type interp_op).
- Local Notation interp_retr := (@interp_retr base_type_code op interp_base_type interp_op).
-
- Lemma interpf_retr_id {t} (e : @exprf interp_base_type t)
- : interpf_retr (fun _ x => x) (fun _ x => x) e = interpf interp_op e.
- Proof using Type.
- induction e; simpl; cbv [LetIn.Let_In]; rewrite_hyp ?*; rewrite ?SmartVarfMap_id; reflexivity.
- Qed.
- Lemma interp_retr_id {t} (e : @expr interp_base_type t)
- : forall x,
- interp_retr (fun _ x => x) (fun _ x => x) e x = interp interp_op e x.
- Proof using Type.
- destruct e; simpl; intros; rewrite interpf_retr_id, SmartVarfMap_id; reflexivity.
- Qed.
-
- Section with_var2.
- Context {var1 var2 : base_type_code -> Type}
- (var1_of_interp : forall t, interp_base_type t -> var1 t)
- (interp_of_var1 : forall t, var1 t -> interp_base_type t)
- (var2_of_interp : forall t, interp_base_type t -> var2 t)
- (interp_of_var2 : forall t, var2 t -> interp_base_type t)
- (interp_of_var12 : forall t x, interp_of_var1 t (var1_of_interp t x)
- = interp_of_var2 t (var2_of_interp t x)).
- Hint Rewrite @flatten_binding_list_SmartVarfMap @List.in_map_iff @List.in_app_iff.
- Lemma interp_of_var12_SmartVarfMap
- t1 e1 t x1 x2
- (H : List.In (existT _ t (x1, x2))
- (flatten_binding_list
- (SmartVarfMap (t:=t1) var1_of_interp e1)
- (SmartVarfMap var2_of_interp e1)))
- : interp_of_var1 t x1 = interp_of_var2 t x2.
- Proof using interp_of_var12.
- repeat first [ progress repeat autorewrite with core in *
- | progress subst
- | progress inversion_sigma
- | progress inversion_prod
- | progress simpl in *
- | progress destruct_head' ex
- | progress destruct_head' and
- | progress destruct_head' or
- | progress destruct_head' sigT
- | progress destruct_head' prod
- | progress rewrite_hyp !*
- | solve [ auto ] ].
- do 2 apply f_equal.
- eapply interp_flat_type_rel_pointwise_flatten_binding_list with (R':=fun _ => eq); [ eassumption | ].
- apply lift_interp_flat_type_rel_pointwise_f_eq; reflexivity.
- Qed.
- Local Hint Resolve List.in_app_or interp_of_var12_SmartVarfMap.
-
- Lemma wff_interpf_retr G {t} (e1 : @exprf var1 t) (e2 : @exprf var2 t)
- (HG : forall t x1 x2,
- List.In (existT _ t (x1, x2)) G
- -> interp_of_var1 t x1 = interp_of_var2 t x2)
- (Hwf : wff G e1 e2)
- : interpf_retr var1_of_interp interp_of_var1 e1
- = interpf_retr var2_of_interp interp_of_var2 e2.
- Proof using interp_of_var12.
- induction Hwf; simpl; rewrite_hyp ?*; try reflexivity; auto.
- { match goal with H : _ |- _ => apply H end.
- intros ???; rewrite List.in_app_iff.
- intros [?|?]; eauto. }
- Qed.
- Lemma wf_interp_retr {t} (e1 : @expr var1 t) (e2 : @expr var2 t)
- (Hwf : wf e1 e2)
- : forall x,
- interp_retr var1_of_interp interp_of_var1 e1 x
- = interp_retr var2_of_interp interp_of_var2 e2 x.
- Proof using interp_of_var12.
- destruct Hwf; simpl; repeat intro; subst; eapply wff_interpf_retr; eauto.
- Qed.
- End with_var2.
-
- Section with_var.
- Context {var : base_type_code -> Type}
- (var_of_interp : forall t, interp_base_type t -> var t)
- (interp_of_var : forall t, var t -> interp_base_type t)
- (var_is_retract : forall t x, interp_of_var t (var_of_interp t x) = x).
- Lemma wff_interpf_retr_correct G {t} (e1 : @exprf var t) (e2 : @exprf interp_base_type t)
- (HG : forall t x1 x2,
- List.In (existT _ t (x1, x2)) G
- -> interp_of_var t x1 = x2)
- (Hwf : wff G e1 e2)
- : interpf_retr var_of_interp interp_of_var e1 = interpf interp_op e2.
- Proof using var_is_retract.
- erewrite wff_interpf_retr, interpf_retr_id; eauto.
- Qed.
- Lemma wf_interp_retr_correct {t} (e1 : @expr var t) (e2 : @expr interp_base_type t)
- (Hwf : wf e1 e2)
- x
- : interp_retr var_of_interp interp_of_var e1 x
- = interp interp_op e2 x.
- Proof using var_is_retract.
- erewrite wf_interp_retr, interp_retr_id; eauto.
- Qed.
- End with_var.
-End language.
diff --git a/src/Compilers/InterpProofs.v b/src/Compilers/InterpProofs.v
deleted file mode 100644
index 5ea1a99a4..000000000
--- a/src/Compilers/InterpProofs.v
+++ /dev/null
@@ -1,66 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Sigma Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-
-Local Open Scope ctype_scope.
-Section language.
- Context (base_type_code : Type).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Context (interp_base_type : base_type_code -> Type).
- Context (op : flat_type (* input tuple *) -> flat_type (* output type *) -> Type).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Context (interp_op : forall src dst, op src dst -> interp_flat_type src -> interp_flat_type dst).
-
- Lemma interpf_LetIn tx ex tC eC
- : Syntax.interpf interp_op (LetIn (tx:=tx) ex (tC:=tC) eC)
- = dlet x := Syntax.interpf interp_op ex in
- Syntax.interpf interp_op (eC x).
- Proof using Type. reflexivity. Qed.
-
- Lemma interpf_SmartVarf t v
- : Syntax.interpf interp_op (SmartVarf (t:=t) v) = v.
- Proof using Type.
- unfold SmartVarf; induction t;
- repeat match goal with
- | _ => reflexivity
- | _ => progress simpl in *
- | _ => progress rewrite_hyp *
- | _ => rewrite <- surjective_pairing
- end.
- Qed.
-
- Lemma interpf_SmartVarVarf {t t'} v x x'
- (Hin : List.In
- (existT (fun t : base_type_code => (exprf base_type_code op (Tbase t) * interp_base_type t)%type)
- t (x, x'))
- (flatten_binding_list (t := t') (SmartVarVarf v) v))
- : interpf interp_op x = x'.
- Proof using Type.
- clear -Hin.
- induction t'; simpl in *; try tauto.
- { intuition (inversion_sigma; inversion_prod; subst; eauto). }
- { apply List.in_app_iff in Hin.
- intuition (inversion_sigma; inversion_prod; subst; eauto). }
- Qed.
-
- Lemma interpf_SmartVarVarf_eq {t t'} v v' x x'
- (Heq : v = v')
- (Hin : List.In
- (existT (fun t : base_type_code => (exprf base_type_code op (Tbase t) * interp_base_type t)%type)
- t (x, x'))
- (flatten_binding_list (t := t') (SmartVarVarf v') v))
- : interpf interp_op x = x'.
- Proof using Type.
- subst; eapply interpf_SmartVarVarf; eassumption.
- Qed.
-End language.
-
-Hint Rewrite @interpf_LetIn @interpf_SmartVarf : reflective_interp.
-Hint Rewrite @interpf_SmartVarVarf using assumption : reflective_interp.
diff --git a/src/Compilers/InterpRewriting.v b/src/Compilers/InterpRewriting.v
deleted file mode 100644
index 048452a35..000000000
--- a/src/Compilers/InterpRewriting.v
+++ /dev/null
@@ -1,197 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Util.Tactics.CacheTerm.
-Require Import Crypto.Util.Tactics.Head.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.NatUtil.
-Require Import Crypto.Util.Tactics.Not.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Section lang.
- Context {base_type}
- {op : flat_type base_type -> flat_type base_type -> Type}
- {interp_base_type : base_type -> Type}
- {interp_op : forall s d, op s d
- -> interp_flat_type interp_base_type s
- -> interp_flat_type interp_base_type d}.
- Local Notation Expr := (@Expr base_type op).
- Local Notation Interp := (@Interp base_type interp_base_type op interp_op).
-
- Definition packaged_expr_functionP A :=
- (fun F : Expr A -> Expr A
- => forall e',
- Wf e'
- -> Wf (F e')
- /\ forall v, Interp (F e') v = Interp e' v).
- Local Notation packaged_expr_function A :=
- (sig (packaged_expr_functionP A)).
-
- Definition compose {A} (f g : packaged_expr_function A)
- : packaged_expr_function A.
- Proof.
- exists (fun x => proj1_sig f (proj1_sig g x)).
- clear.
- abstract (
- destruct f as [f Hf], g as [g Hg]; cbn [proj1_sig];
- intros e' Wfe; split; [ apply Hf, Hg, Wfe | ];
- intro x; etransitivity; [ apply Hf, Hg, Wfe | apply Hg, Wfe ]
- ).
- Defined.
-
- Definition id_package {A} : packaged_expr_function A
- := exist (packaged_expr_functionP A)
- id
- (fun e' Wfe' => conj Wfe' (fun v => eq_refl)).
-
- Inductive reified_transformation :=
- | base (idx : nat)
- | transform (idx : nat) (rest : reified_transformation)
- | cond (test : bool) (iftrue iffalse : reified_transformation).
- Fixpoint denote {A}
- (ls : list (packaged_expr_function A))
- (ls' : list { x : Expr A | Wf x })
- default
- (f : reified_transformation)
- := match f with
- | base idx => proj1_sig (List.nth_default default ls' idx)
- | transform idx rest
- => proj1_sig (List.nth_default id_package ls idx)
- (denote ls ls' default rest)
- | cond test iftrue iffalse
- => if test
- then denote ls ls' default iftrue
- else denote ls ls' default iffalse
- end.
- Fixpoint reduce (f : reified_transformation) : reified_transformation
- := match f with
- | base idx => base idx
- | transform idx rest => reduce rest
- | cond test iftrue iffalse
- => match reduce iftrue, reduce iffalse with
- | base idx0 as t, base idx1 as f
- => if nat_beq idx0 idx1
- then base idx0
- else cond test t f
- | t, f => cond test t f
- end
- end.
- Lemma Wf_denote A ctx es d f : Wf (@denote A ctx es d f).
- Proof.
- induction f; simpl; unfold proj1_sig; break_innermost_match; split_and; auto.
- match goal with H : _ |- _ => apply H; assumption end.
- Qed.
- Lemma Wf_denote_iff_True A ctx es d f : Wf (@denote A ctx es d f) <-> True.
- Proof. split; auto using Wf_denote. Qed.
- Lemma Interp_denote_reduce A ctx es d f
- : forall v, Interp (@denote A ctx es d f) v = Interp (@denote A nil es d (reduce f)) v.
- Proof.
- induction f; simpl; unfold proj1_sig; break_innermost_match;
- nat_beq_to_eq; subst;
- try reflexivity; auto.
- intro; rewrite <- IHf.
- match goal with H : _ |- _ => apply H, Wf_denote end.
- Qed.
-End lang.
-
-Local Ltac find ctx f :=
- lazymatch ctx with
- | (exist _ f _ :: _)%list => constr:(0)
- | (_ :: ?ctx)%list
- => let v := find ctx f in
- constr:(S v)
- end.
-
-Local Ltac reify_transformation interp_base_type interp_op ctx es T cont :=
- let reify_transformation := reify_transformation interp_base_type interp_op in
- let ExprA := type of T in
- let packageP := lazymatch type of T with
- | @Expr ?base_type_code ?op ?A
- => constr:(@packaged_expr_functionP base_type_code op interp_base_type interp_op A)
- end in
- let es := lazymatch es with
- | tt => constr:(@nil { x : ExprA | Wf x })
- | _ => es
- end in
- let ctx := lazymatch ctx with
- | tt => constr:(@nil (sig packageP))
- | _ => ctx
- end in
- lazymatch T with
- | ?f ?e
- => let ctx := lazymatch ctx with
- | context[exist _ f _] => ctx
- | _ => let hf := head f in
- let fId := fresh hf in
- let rfPf :=
- cache_proof_with_type_by
- (packageP f)
- ltac:(refine (fun e Hwf
- => (fun Hwf'
- => conj Hwf' (fun v => _)) _);
- [ autorewrite with reflective_interp; reflexivity
- | auto with wf ])
- fId in
- constr:(cons (exist packageP f rfPf)
- ctx)
- end in
- reify_transformation
- ctx es e
- ltac:(fun ctx es re
- => let idx := find ctx f in
- cont ctx es (transform idx re))
- | match ?b with true => ?t | false => ?f end
- => reify_transformation
- ctx es t
- ltac:(fun ctx es rt
- => reify_transformation
- ctx es f
- ltac:(fun ctx es rf
- => reify_transformation
- ctx es t
- ltac:(fun ctx es rt
- => cont ctx es (cond b rt rf))))
- | _ => let es := lazymatch es with
- | context[exist _ T _] => es
- | _
- => let Hwf := lazymatch goal with
- | [ Hwf : Wf T |- _ ] => Hwf
-
- | _
- => let Hwf := fresh "Hwf" in
- cache_proof_with_type_by
- (Wf T)
- ltac:(idtac; solve_wf_side_condition)
- Hwf
- end in
- constr:(cons (exist Wf T Hwf) es)
- end in
- let idx := find es T in
- cont ctx es (base idx)
- end.
-Ltac finish_rewrite_reflective_interp_cached :=
- rewrite ?Wf_denote_iff_True;
- cbv [reduce nat_beq];
- try (rewrite Interp_denote_reduce;
- cbv [reduce nat_beq];
- cbv [denote List.nth_default List.nth_error];
- cbn [proj1_sig]).
-Ltac rewrite_reflective_interp_cached_then ctx es cont :=
- let e := match goal with
- | [ |- context[@Interp _ _ _ _ _ ?e] ]
- => let test := match goal with _ => not is_var e end in
- e
- end in
- lazymatch goal with
- | [ |- context[@Interp ?base_type ?interp_base_type ?op ?interp_op _ e] ]
- => reify_transformation
- interp_base_type interp_op ctx es e
- ltac:(fun ctx es r
- => lazymatch es with
- | cons ?default _
- => change e with (denote ctx es default r)
- end;
- finish_rewrite_reflective_interp_cached;
- cont ctx es)
- end.
-Ltac rewrite_reflective_interp_cached :=
- rewrite_reflective_interp_cached_then tt tt ltac:(fun _ _ => idtac).
diff --git a/src/Compilers/InterpSideConditions.v b/src/Compilers/InterpSideConditions.v
deleted file mode 100644
index 42467e7fc..000000000
--- a/src/Compilers/InterpSideConditions.v
+++ /dev/null
@@ -1,40 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Notations.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d)
- (interped_op_side_conditions : forall s d, op s d -> interp_flat_type interp_base_type s -> pointed_Prop).
-
- Local Notation exprf := (@exprf base_type_code op interp_base_type).
- Local Notation expr := (@expr base_type_code op interp_base_type).
- Local Notation Expr := (@Expr base_type_code op).
-
- Fixpoint interpf_side_conditions_gen {t} (e : exprf t) : pointed_Prop * interp_flat_type interp_base_type t
- := match e with
- | TT => (trivial, tt)
- | Var t v => (trivial, v)
- | Op t1 tR opc args
- => let '(args_cond, argsv) := @interpf_side_conditions_gen _ args in
- (args_cond /\ interped_op_side_conditions _ _ opc argsv, interp_op _ _ opc argsv)
- | LetIn tx ex tC eC
- => let '(x_cond, xv) := @interpf_side_conditions_gen _ ex in
- let '(c_cond, cv) := @interpf_side_conditions_gen _ (eC xv) in
- (x_cond /\ c_cond, cv)
- | Pair tx ex ty ey
- => let '(x_cond, xv) := @interpf_side_conditions_gen _ ex in
- let '(y_cond, yv) := @interpf_side_conditions_gen _ ey in
- (x_cond /\ y_cond, (xv, yv))
- end%pointed_prop.
- Definition interpf_side_conditions {t} e : pointed_Prop
- := fst (@interpf_side_conditions_gen t e).
- Definition interp_side_conditions {t} (e : expr t) : interp_flat_type interp_base_type (domain t) -> pointed_Prop
- := fun x => interpf_side_conditions (invert_Abs e x).
- Definition InterpSideConditions {t} (e : Expr t) : interp_flat_type interp_base_type (domain t) -> pointed_Prop
- := interp_side_conditions (e _).
-End language.
diff --git a/src/Compilers/InterpWf.v b/src/Compilers/InterpWf.v
deleted file mode 100644
index e1572ceed..000000000
--- a/src/Compilers/InterpWf.v
+++ /dev/null
@@ -1,80 +0,0 @@
-Require Import Coq.Strings.String Coq.Classes.RelationClasses.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Util.Tuple.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Notations.
-Local Open Scope ctype_scope.
-Local Open Scope expr_scope.
-
-Section language.
- Context {base_type_code : Type}
- {interp_base_type : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst).
-
- Local Notation exprf := (@exprf base_type_code op interp_base_type).
- Local Notation expr := (@expr base_type_code op interp_base_type).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation interpf := (@interpf base_type_code interp_base_type op interp_op).
- Local Notation interp := (@interp base_type_code interp_base_type op interp_op).
- Local Notation Interp := (@Interp base_type_code interp_base_type op interp_op).
-
- Lemma eq_in_flatten_binding_list
- {t x x' T e}
- (HIn : List.In (existT (fun t : base_type_code => (interp_base_type t * interp_base_type t)%type) t (x, x')%core)
- (flatten_binding_list (t:=T) e e))
- : x = x'.
- Proof using Type.
- induction T; simpl in *; [ | | rewrite List.in_app_iff in HIn ];
- repeat first [ progress destruct_head or
- | progress destruct_head False
- | progress destruct_head and
- | progress inversion_sigma
- | progress inversion_prod
- | progress subst
- | solve [ eauto ] ].
- Qed.
-
-
- Local Hint Resolve List.in_app_or List.in_or_app eq_in_flatten_binding_list.
-
- Section wf.
- Lemma interpf_wff
- {t} {e1 e2 : exprf t}
- {G}
- (HG : forall t x x',
- List.In (existT (fun t : base_type_code => (interp_base_type t * interp_base_type t)%type) t (x, x')%core) G
- -> x = x')
- (Rwf : wff G e1 e2)
- : interpf e1 = interpf e2.
- Proof using Type.
- induction Rwf; simpl; auto;
- specialize_by auto; try congruence.
- rewrite_hyp !*; auto.
- repeat match goal with
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => setoid_rewrite List.in_app_iff in H
- end.
- match goal with
- | [ H : _ |- _ ]
- => apply H; intros; destruct_head' or; solve [ eauto ]
- end.
- Qed.
-
- Local Hint Resolve interpf_wff.
-
- Lemma interp_wf
- {t} {e1 e2 : expr t}
- (Rwf : wf e1 e2)
- : forall x, interp e1 x = interp e2 x.
- Proof using Type.
- destruct Rwf; simpl; eauto.
- Qed.
- End wf.
-End language.
diff --git a/src/Compilers/InterpWfRel.v b/src/Compilers/InterpWfRel.v
deleted file mode 100644
index 46be4220d..000000000
--- a/src/Compilers/InterpWfRel.v
+++ /dev/null
@@ -1,94 +0,0 @@
-Require Import Coq.Strings.String Coq.Classes.RelationClasses.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Util.Tuple.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Notations.
-Local Open Scope ctype_scope.
-Local Open Scope expr_scope.
-
-Section language.
- Context {base_type_code : Type}
- {interp_base_type1 interp_base_type2 : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (interp_op1 : forall src dst, op src dst -> interp_flat_type interp_base_type1 src -> interp_flat_type interp_base_type1 dst)
- (interp_op2 : forall src dst, op src dst -> interp_flat_type interp_base_type2 src -> interp_flat_type interp_base_type2 dst)
- {R : forall t, interp_base_type1 t -> interp_base_type2 t -> Prop}
- (Rop : forall src dst op sv1 sv2, interp_flat_type_rel_pointwise R sv1 sv2
- -> interp_flat_type_rel_pointwise
- R (interp_op1 src dst op sv1) (interp_op2 src dst op sv2)).
-
- Local Notation exprf1 := (@exprf base_type_code op interp_base_type1).
- Local Notation exprf2 := (@exprf base_type_code op interp_base_type2).
- Local Notation expr1 := (@expr base_type_code op interp_base_type1).
- Local Notation expr2 := (@expr base_type_code op interp_base_type2).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation interpf1 := (@interpf base_type_code interp_base_type1 op interp_op1).
- Local Notation interpf2 := (@interpf base_type_code interp_base_type2 op interp_op2).
- Local Notation interp1 := (@interp base_type_code interp_base_type1 op interp_op1).
- Local Notation interp2 := (@interp base_type_code interp_base_type2 op interp_op2).
- Local Notation Interp1 := (@Interp base_type_code interp_base_type1 op interp_op1).
- Local Notation Interp2 := (@Interp base_type_code interp_base_type2 op interp_op2).
-
- Lemma interp_flat_type_rel_pointwise_flatten_binding_list
- {t x x' T e1 e2}
- (Hpointwise : interp_flat_type_rel_pointwise R e1 e2)
- (HIn : List.In (existT (fun t : base_type_code => (interp_base_type1 t * interp_base_type2 t)%type) t (x, x')%core)
- (flatten_binding_list (t:=T) e1 e2))
- : R t x x'.
- Proof using Type.
- induction T; simpl in *; try tauto; [ | rewrite List.in_app_iff in HIn ];
- repeat first [ progress destruct_head or
- | progress destruct_head False
- | progress destruct_head and
- | progress inversion_sigma
- | progress inversion_prod
- | progress subst
- | solve [ eauto ] ].
- Qed.
-
- Local Hint Resolve List.in_app_or List.in_or_app interp_flat_type_rel_pointwise_flatten_binding_list.
-
- Section wf.
- Lemma interpf_wff
- {t} {e1 : exprf1 t} {e2 : exprf2 t}
- {G}
- (HG : forall t x x',
- List.In (existT (fun t : base_type_code => (interp_base_type1 t * interp_base_type2 t)%type) t (x, x')%core) G
- -> R t x x')
- (Rwf : wff G e1 e2)
- : interp_flat_type_rel_pointwise R (interpf1 e1) (interpf2 e2).
- Proof using Type*.
- induction Rwf; simpl; auto.
- repeat match goal with
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => setoid_rewrite List.in_app_iff in H
- end.
- match goal with
- | [ H : _ |- _ ]
- => apply H; intros; destruct_head' or; solve [ eauto ]
- end.
- Qed.
-
- Local Hint Resolve interpf_wff.
-
- Lemma interp_wf
- {t} {e1 : expr1 t} {e2 : expr2 t}
- (Rwf : wf e1 e2)
- : interp_type_rel_pointwise R (interp1 e1) (interp2 e2).
- Proof using Type*.
- destruct Rwf; simpl; repeat intro; eauto.
- Qed.
-
- Lemma InterpWf
- {t} {e : Expr t}
- (Rwf : Wf e)
- : interp_type_rel_pointwise R (Interp1 e) (Interp2 e).
- Proof using Type*.
- unfold Interp, Wf in *; apply interp_wf; simpl; intuition.
- Qed.
- End wf.
-End language.
diff --git a/src/Compilers/Intros.v b/src/Compilers/Intros.v
deleted file mode 100644
index e51c9c2fa..000000000
--- a/src/Compilers/Intros.v
+++ /dev/null
@@ -1,94 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-
-Section language.
- Context {base_type_code}
- {var : base_type_code -> Type}.
-
- Section cps.
- Context {rT : Type}
- (Forall : forall T (R : var T -> rT), rT).
-
- Fixpoint intros_interp_flat_type_cps
- {t : flat_type base_type_code}
- : forall (R : interp_flat_type var t -> rT),
- rT
- := match t return (interp_flat_type var t -> rT) -> rT with
- | Tbase T => fun R => Forall _ (fun v : var T => R v)
- | Unit => fun R => R tt
- | Prod A B
- => fun R : interp_flat_type _ A * interp_flat_type _ B -> _
- => @intros_interp_flat_type_cps
- A
- (fun a
- => @intros_interp_flat_type_cps
- B
- (fun b
- => R (a, b)))
- end.
- End cps.
- Definition intros_interp_flat_type_Prop {t}
- := @intros_interp_flat_type_cps Prop (fun T R => forall v : var T, R v) t.
- Definition intros_interp_flat_type_Type {t}
- := @intros_interp_flat_type_cps Type (fun T R => forall v : var T, R v) t.
-
- Fixpoint intros_interp_flat_type
- {t : flat_type base_type_code}
- : forall {R : interp_flat_type var t -> Type},
- @intros_interp_flat_type_Type t R
- -> (forall v : interp_flat_type var t, R v)
- := match t return forall R : interp_flat_type var t -> Type,
- @intros_interp_flat_type_Type t R -> (forall v : interp_flat_type _ t, R v)
- with
- | Tbase T => fun R f => f
- | Unit => fun R p 'tt => p
- | Prod A B
- => fun (R : interp_flat_type _ A * interp_flat_type _ B -> _)
- (f : intros_interp_flat_type_Type
- (fun a => intros_interp_flat_type_Type (fun b => R (a, b))))
- '((a, b) : interp_flat_type _ A * interp_flat_type _ B)
- => @intros_interp_flat_type
- B _ (@intros_interp_flat_type A _ f a) b
- end.
-
- Fixpoint introsP_interp_flat_type
- {t : flat_type base_type_code}
- : forall {R : interp_flat_type var t -> Prop},
- @intros_interp_flat_type_Prop t R
- -> (forall v : interp_flat_type var t, R v)
- := match t return forall R : interp_flat_type var t -> Prop,
- @intros_interp_flat_type_Prop t R -> (forall v : interp_flat_type _ t, R v)
- with
- | Tbase T => fun R f => f
- | Unit => fun R p 'tt => p
- | Prod A B
- => fun (R : interp_flat_type _ A * interp_flat_type _ B -> _)
- (f : intros_interp_flat_type_Prop
- (fun a => intros_interp_flat_type_Prop (fun b => R (a, b))))
- '((a, b) : interp_flat_type _ A * interp_flat_type _ B)
- => @introsP_interp_flat_type
- B _ (@introsP_interp_flat_type A _ f a) b
- end.
-End language.
-
-Ltac post_intro_interp_flat_type_intros :=
- let do_cbv _ :=
- (cbv [intros_interp_flat_type_Type intros_interp_flat_type_Prop intros_interp_flat_type_cps]) in
- lazymatch goal with
- | [ |- @intros_interp_flat_type_Type _ ?var ?t ?R ]
- => let t' := (eval cbv in t) in
- change (@intros_interp_flat_type_Type _ var t' (fun v => id (R v)));
- do_cbv (); post_intro_interp_flat_type_intros
- | [ |- @intros_interp_flat_type_Prop _ ?var ?t ?R ]
- => let t' := (eval cbv in t) in
- change (@intros_interp_flat_type_Prop _ var t' (fun v => id (R v)));
- do_cbv (); post_intro_interp_flat_type_intros
- | [ |- forall x, _ ] => intro; post_intro_interp_flat_type_intros
- | [ |- id ?P ] => change P
- end.
-Ltac intro_interp_flat_type :=
- lazymatch goal with
- | [ |- forall v : interp_flat_type ?var ?t, @?R v ]
- => let t' := (eval compute in t) in
- refine (@intros_interp_flat_type _ var t' (fun v => id (R v)) _);
- post_intro_interp_flat_type_intros
- end.
diff --git a/src/Compilers/Linearize.v b/src/Compilers/Linearize.v
deleted file mode 100644
index e5d136f5f..000000000
--- a/src/Compilers/Linearize.v
+++ /dev/null
@@ -1,97 +0,0 @@
-(** * Linearize: Place all and only operations in let binders *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-(*Require Import Crypto.Util.Tactics.*)
-
-Local Open Scope ctype_scope.
-Section language.
- Context (let_bind_op_args : bool)
- {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation Tbase := (@Tbase base_type_code).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var : base_type_code -> Type}.
- Local Notation exprf := (@exprf base_type_code op var).
- Local Notation expr := (@expr base_type_code op var).
-
- Section under_lets.
- Fixpoint under_letsf' (bind_pairs : bool) {t} (e : exprf t)
- : forall {tC} (C : interp_flat_type var t + exprf t -> exprf tC), exprf tC
- := match e in Syntax.exprf _ _ t return forall {tC} (C : interp_flat_type var t + exprf t -> exprf tC), exprf tC with
- | LetIn _ ex _ eC
- => fun _ C => @under_letsf' true _ ex _ (fun v =>
- match v with
- | inl v => @under_letsf' false _ (eC v) _ C
- | inr v => LetIn v (fun v => @under_letsf' false _ (eC v) _ C)
- end)
- | TT => fun _ C => C (inl tt)
- | Var _ x => fun _ C => C (inl x)
- | Op _ _ op args as e
- => if let_bind_op_args
- then fun _ C => LetIn e (fun v => C (inl v))
- else fun _ C => C (inr e)
- | Pair A x B y => fun _ C => @under_letsf' bind_pairs A x _ (fun x =>
- @under_letsf' bind_pairs B y _ (fun y =>
- match x, y, bind_pairs with
- | inl x, inl y, _ => C (inl (x, y))
- | inl x, inr y, false => C (inr (Pair (SmartVarf x) y))
- | inr x, inl y, false => C (inr (Pair x (SmartVarf y)))
- | inr x, inr y, false => C (inr (Pair x y))
- | inl x, inr y, true => LetIn y (fun y => C (inl (x, y)))
- | inr x, inl y, true => LetIn x (fun x => C (inl (x, y)))
- | inr x, inr y, true => LetIn x (fun x => LetIn y (fun y => C (inl (x, y))))
- end))
- end.
- Definition under_letsf {t} (e : exprf t)
- {tC} (C : exprf t -> exprf tC)
- : exprf tC
- := under_letsf' false e (fun v => match v with inl v => C (SmartVarf v) | inr v => C v end).
- End under_lets.
-
- Fixpoint linearizef_gen {t} (e : exprf t) : exprf t
- := match e in Syntax.exprf _ _ t return exprf t with
- | LetIn _ ex _ eC
- => under_letsf (LetIn (@linearizef_gen _ ex) (fun x => @linearizef_gen _ (eC x))) (fun x => x)
- | TT => TT
- | Var _ x => Var x
- | Op _ _ op args
- => if let_bind_op_args
- then under_letsf (@linearizef_gen _ args) (fun args => LetIn (Op op args) SmartVarf)
- else under_letsf (@linearizef_gen _ args) (fun args => Op op args)
- | Pair A ex B ey
- => under_letsf (Pair (@linearizef_gen _ ex) (@linearizef_gen _ ey)) (fun v => v)
- end.
-
- Definition linearize_gen {t} (e : expr t) : expr t
- := match e in Syntax.expr _ _ t return expr t with
- | Abs _ _ f => Abs (fun x => linearizef_gen (f x))
- end.
- End with_var.
-
- Definition Linearize_gen {t} (e : Expr t) : Expr t
- := fun var => linearize_gen (e _).
-End language.
-
-Global Arguments under_letsf' _ {_ _ _} bind_pairs {_} _ {tC} _.
-Global Arguments under_letsf _ {_ _ _ _} _ {tC} _.
-Global Arguments linearizef_gen _ {_ _ _ _} _.
-Global Arguments linearize_gen _ {_ _ _ _} _.
-Global Arguments Linearize_gen _ {_ _ _} _ var.
-
-Definition linearizef {base_type_code op var t} e
- := @linearizef_gen false base_type_code op var t e.
-Definition a_normalf {base_type_code op var t} e
- := @linearizef_gen true base_type_code op var t e.
-Definition linearize {base_type_code op var t} e
- := @linearize_gen false base_type_code op var t e.
-Definition a_normal {base_type_code op var t} e
- := @linearize_gen true base_type_code op var t e.
-Definition Linearize {base_type_code op t} e
- := @Linearize_gen false base_type_code op t e.
-Definition ANormal {base_type_code op t} e
- := @Linearize_gen true base_type_code op t e.
diff --git a/src/Compilers/LinearizeInterp.v b/src/Compilers/LinearizeInterp.v
deleted file mode 100644
index 160826b70..000000000
--- a/src/Compilers/LinearizeInterp.v
+++ /dev/null
@@ -1,136 +0,0 @@
-(** * Linearize: Place all and only operations in let binders *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.InterpProofs.
-Require Import Crypto.Compilers.Linearize.
-Require Import Crypto.Util.Sigma Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {interp_base_type : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Local Hint Extern 1 => eapply interpf_SmartVarVarf.
-
- Local Ltac t_fin :=
- repeat match goal with
- | _ => reflexivity
- | _ => progress unfold LetIn.Let_In
- | _ => progress simpl in *
- | _ => progress intros
- | _ => progress inversion_sigma
- | _ => progress inversion_prod
- | _ => solve [ intuition eauto | eauto | symmetry; eauto ]
- | _ => apply (f_equal (interp_op _ _ _))
- | _ => apply (f_equal2 (@pair _ _))
- | _ => progress specialize_by assumption
- | _ => progress subst
- | [ H : context[List.In _ (_ ++ _)] |- _ ] => setoid_rewrite List.in_app_iff in H
- | [ H : or _ _ |- _ ] => destruct H
- | _ => progress break_match
- | _ => rewrite <- !surjective_pairing
- | [ H : ?x = _, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : _ |- _ ] => apply H
- | [ H : _ |- _ ] => rewrite H
- | [ H : _ |- _ ] => erewrite H by reflexivity
- | _ => rewrite interpf_SmartVarf
- end.
-
- Lemma interpf_under_letsf' let_bind_op_args bind_pairs {t tC} (ex : exprf t) (eC : _ -> exprf tC)
- (eC_resp : forall x y,
- x = interpf interp_op y
- -> interpf interp_op (eC (inr y)) = interpf interp_op (eC (inl x)))
- : interpf interp_op (under_letsf' let_bind_op_args bind_pairs ex eC)
- = let x := interpf interp_op ex in interpf interp_op (eC (inl x)).
- Proof using Type.
- clear -eC_resp; revert bind_pairs.
- induction ex; t_fin.
- Qed.
-
- Lemma interpf_under_letsf let_bind_op_args {t tC} (ex : exprf t) (eC : _ -> exprf tC)
- (eC_resp : forall x y,
- interpf interp_op x = interpf interp_op y
- -> interpf interp_op (eC x) = interpf interp_op (eC y))
- : interpf interp_op (under_letsf let_bind_op_args ex eC)
- = let x := interpf interp_op ex in interpf interp_op (eC (SmartMap.SmartVarf x)).
- Proof using Type.
- unfold under_letsf; rewrite interpf_under_letsf'; t_fin.
- Qed.
-
- Section gen.
- Context (let_bind_op_args : bool).
-
- Lemma interpf_linearizef_gen {t} e
- : interpf interp_op (linearizef_gen let_bind_op_args (t:=t) e) = interpf interp_op e.
- Proof using Type.
- clear.
- induction e;
- repeat first [ progress rewrite ?interpf_under_letsf, ?interpf_SmartVarf
- | progress simpl
- | t_fin ].
- Qed.
-
- Local Hint Resolve interpf_linearizef_gen.
-
- Lemma interp_linearize_gen {t} e
- : forall x, interp interp_op (linearize_gen let_bind_op_args (t:=t) e) x = interp interp_op e x.
- Proof using Type.
- induction e; simpl; eauto.
- Qed.
-
- Lemma InterpLinearize_gen {t} (e : Expr t)
- : forall x, Interp interp_op (Linearize_gen let_bind_op_args e) x = Interp interp_op e x.
- Proof using Type.
- unfold Interp, Linearize_gen.
- eapply interp_linearize_gen.
- Qed.
-
- Lemma InterpLinearize_gen_ind {t} (P : _ -> Prop) {e : Expr t} {x}
- : P (Interp interp_op e x) -> P (Interp interp_op (Linearize_gen let_bind_op_args e) x).
- Proof using Type. rewrite InterpLinearize_gen; exact id. Qed.
- End gen.
-
- Definition interpf_linearizef {t} e
- : interpf interp_op (linearizef (t:=t) e) = interpf interp_op e
- := interpf_linearizef_gen _ e.
- Definition interpf_a_normalf {t} e
- : interpf interp_op (a_normalf (t:=t) e) = interpf interp_op e
- := interpf_linearizef_gen _ e.
-
- Definition interp_linearize {t} e
- : forall x, interp interp_op (linearize (t:=t) e) x = interp interp_op e x
- := interp_linearize_gen _ e.
- Definition interp_a_normal {t} e
- : forall x, interp interp_op (a_normal (t:=t) e) x = interp interp_op e x
- := interp_linearize_gen _ e.
-
- Definition InterpLinearize {t} (e : Expr t)
- : forall x, Interp interp_op (Linearize e) x = Interp interp_op e x
- := InterpLinearize_gen _ e.
- Definition InterpANormal {t} (e : Expr t)
- : forall x, Interp interp_op (ANormal e) x = Interp interp_op e x
- := InterpLinearize_gen _ e.
-
- Definition InterpLinearize_ind {t} (P : _ -> Prop) {e : Expr t} {x}
- : P (Interp interp_op e x) -> P (Interp interp_op (Linearize e) x)
- := InterpLinearize_gen_ind _ P.
- Definition InterpANormal_ind {t} (P : _ -> Prop) {e : Expr t} {x}
- : P (Interp interp_op e x) -> P (Interp interp_op (ANormal e) x)
- := InterpLinearize_gen_ind _ P.
-End language.
-
-Hint Rewrite @interpf_under_letsf : reflective_interp.
-Hint Rewrite @InterpLinearize_gen @interp_linearize_gen @interpf_linearizef_gen @InterpLinearize @interp_linearize @interpf_linearizef @InterpANormal @interp_a_normal @interpf_a_normalf : reflective_interp.
diff --git a/src/Compilers/LinearizeWf.v b/src/Compilers/LinearizeWf.v
deleted file mode 100644
index d2e7f2807..000000000
--- a/src/Compilers/LinearizeWf.v
+++ /dev/null
@@ -1,179 +0,0 @@
-(** * Linearize: Place all and only operations in let binders *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.Linearize.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation Tbase := (@Tbase base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}.
-
- Local Ltac t_fin_step tac :=
- match goal with
- | _ => assumption
- | _ => progress simpl in *
- | _ => progress subst
- | _ => progress inversion_sigma
- | _ => progress destruct_head'_sig
- | _ => progress destruct_head'_and
- | _ => setoid_rewrite List.in_app_iff
- | [ H : context[List.In _ (_ ++ _)] |- _ ] => setoid_rewrite List.in_app_iff in H
- | _ => progress intros
- | _ => solve [ eauto ]
- | _ => solve [ intuition (subst; eauto) ]
- | [ H : forall (x : prod _ _) (y : prod _ _), _ |- _ ] => specialize (fun x x' y y' => H (x, x') (y, y'))
- | _ => rewrite !List.app_assoc
- | [ H : _ \/ _ |- _ ] => destruct H
- | [ H : _ |- _ ] => apply H
- | [ H : forall G : list _, _ |- _ ] => apply (H nil)
- | _ => eapply wff_in_impl_Proper; [ solve [ eauto using wff_SmartVarf ] | solve [ repeat t_fin_step tac ] ]
- | _ => progress tac
- | [ |- wff _ _ _ ] => constructor
- | [ |- wf _ _ _ ] => constructor
- | _ => break_innermost_match_step
- | _ => progress inversion_expr
- | _ => congruence
- | _ => progress destruct_head' sum
- | _ => progress unfold invert_Op in *
- | _ => break_innermost_match_hyps_step
- end.
- Local Ltac t_fin tac := repeat t_fin_step tac.
-
- Local Hint Constructors Wf.wff.
- Local Hint Resolve List.in_app_or List.in_or_app.
-
- Section gen1.
- Context (let_bind_op_args : bool).
-
- Fixpoint wff_under_letsf' G {t} e1 e2 {tC} eC1 eC2
- let_bind_pairs
- (wf : @wff var1 var2 G t e1 e2)
- (H : forall (x1 : interp_flat_type var1 t) (x2 : interp_flat_type var2 t),
- wff (flatten_binding_list x1 x2 ++ G) (eC1 (inl x1)) (eC2 (inl x2)))
- (H' : forall G' (x y : exprf t),
- wff (G' ++ G) x y
- -> wff (G' ++ G) (eC1 (inr x)) (eC2 (inr y)))
- {struct e1}
- : @wff var1 var2 G tC (under_letsf' let_bind_op_args let_bind_pairs e1 eC1) (under_letsf' let_bind_op_args let_bind_pairs e2 eC2).
- Proof using Type.
- revert H.
- set (e1v := e1) in *.
- destruct e1 as [ | | ? ? ? args | tx ex tC0 eC0 | ? ex ? ey ];
- [ clear wff_under_letsf'
- | clear wff_under_letsf'
- | clear wff_under_letsf'
- | generalize (fun G => match e1v return match e1v with LetIn _ _ _ _ => _ | _ => _ end with
- | LetIn _ ex _ eC => wff_under_letsf' G _ ex
- | _ => I
- end);
- generalize (fun G => match e1v return match e1v with LetIn tx0 _ tC1 e0 => _ | _ => _ end with
- | LetIn _ ex tC' eC => fun x => wff_under_letsf' G tC' (eC x)
- | _ => I
- end);
- clear wff_under_letsf'
- | generalize (fun G => match e1v return match e1v with Pair _ _ _ _ => _ | _ => _ end with
- | Pair _ ex _ ey => wff_under_letsf' G _ ex
- | _ => I
- end);
- generalize (fun G => match e1v return match e1v with Pair _ _ _ _ => _ | _ => _ end with
- | Pair _ ex _ ey => wff_under_letsf' G _ ey
- | _ => I
- end);
- clear wff_under_letsf' ];
- subst e1v;
- cbv beta iota;
- (invert_one_expr e2 || destruct e2); intros; try break_innermost_match_step; try exact I; intros;
- inversion_wf;
- t_fin idtac.
- Qed.
-
- Lemma wff_under_letsf G {t} e1 e2 {tC} eC1 eC2
- (wf : @wff var1 var2 G t e1 e2)
- (H' : forall G' (x y : exprf t),
- wff (G' ++ G) x y
- -> wff (G' ++ G) (eC1 x) (eC2 y))
- : @wff var1 var2 G tC (under_letsf let_bind_op_args e1 eC1) (under_letsf let_bind_op_args e2 eC2).
- Proof using Type.
- apply wff_under_letsf'; t_fin idtac.
- Qed.
- End gen1.
-
- Local Hint Resolve wff_under_letsf.
- Local Hint Constructors or.
- Local Hint Extern 1 => progress unfold List.In in *.
- Local Hint Resolve wff_in_impl_Proper.
- Local Hint Resolve wff_SmartVarf.
-
- Section gen2.
- Context (let_bind_op_args : bool).
-
- Lemma wff_linearizef_gen G {t} e1 e2
- : @wff var1 var2 G t e1 e2
- -> @wff var1 var2 G t (linearizef_gen let_bind_op_args e1) (linearizef_gen let_bind_op_args e2).
- Proof using Type.
- induction 1; t_fin ltac:(apply wff_under_letsf).
- Qed.
-
- Local Hint Resolve wff_linearizef_gen.
-
- Lemma wf_linearize_gen {t} e1 e2
- : @wf var1 var2 t e1 e2
- -> @wf var1 var2 t (linearize_gen let_bind_op_args e1) (linearize_gen let_bind_op_args e2).
- Proof using Type.
- destruct 1; constructor; auto.
- Qed.
- End gen2.
- End with_var.
-
- Section gen.
- Context (let_bind_op_args : bool).
-
- Lemma Wf_Linearize_gen {t} (e : Expr t) : Wf e -> Wf (Linearize_gen let_bind_op_args e).
- Proof using Type.
- intros wf var1 var2; apply wf_linearize_gen, wf.
- Qed.
- End gen.
-
- Definition wff_linearizef {var1 var2} G {t} e1 e2
- : @wff var1 var2 G t e1 e2
- -> @wff var1 var2 G t (linearizef e1) (linearizef e2)
- := wff_linearizef_gen _ G e1 e2.
- Definition wff_a_normalf {var1 var2} G {t} e1 e2
- : @wff var1 var2 G t e1 e2
- -> @wff var1 var2 G t (a_normalf e1) (a_normalf e2)
- := wff_linearizef_gen _ G e1 e2.
-
- Definition wf_linearize {var1 var2 t} e1 e2
- : @wf var1 var2 t e1 e2
- -> @wf var1 var2 t (linearize e1) (linearize e2)
- := wf_linearize_gen _ e1 e2.
- Definition wf_a_normal {var1 var2 t} e1 e2
- : @wf var1 var2 t e1 e2
- -> @wf var1 var2 t (a_normal e1) (a_normal e2)
- := wf_linearize_gen _ e1 e2.
-
- Definition Wf_Linearize {t} (e : Expr t) : Wf e -> Wf (Linearize e)
- := Wf_Linearize_gen _ e.
- Definition Wf_ANormal {t} (e : Expr t) : Wf e -> Wf (ANormal e)
- := Wf_Linearize_gen _ e.
-End language.
-
-Hint Resolve Wf_Linearize_gen Wf_Linearize Wf_ANormal : wf.
diff --git a/src/Compilers/Map.v b/src/Compilers/Map.v
deleted file mode 100644
index 9fe9f7011..000000000
--- a/src/Compilers/Map.v
+++ /dev/null
@@ -1,30 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {var1 var2 : base_type_code -> Type}
- (fvar12 : forall t, var1 t -> var2 t)
- (fvar21 : forall t, var2 t -> var1 t).
- Local Notation exprf := (@exprf base_type_code op).
- Fixpoint mapf_interp_flat_type {t} (e : interp_flat_type var2 t) {struct t}
- : interp_flat_type var1 t
- := match t return interp_flat_type _ t -> interp_flat_type _ t with
- | Tbase _ => fvar21 _
- | Unit => fun v : unit => v
- | Prod x y => fun xy => (@mapf_interp_flat_type _ (fst xy),
- @mapf_interp_flat_type _ (snd xy))
- end e.
-
- Fixpoint mapf {t} (e : @exprf var1 t) : @exprf var2 t
- := match e in Syntax.exprf _ _ t return exprf t with
- | TT => TT
- | Var _ x => Var (fvar12 _ x)
- | Op _ _ op args => Op op (@mapf _ args)
- | LetIn _ ex _ eC => LetIn (@mapf _ ex) (fun x => @mapf _ (eC (mapf_interp_flat_type x)))
- | Pair _ ex _ ey => Pair (@mapf _ ex) (@mapf _ ey)
- end.
-End language.
-
-Global Arguments mapf_interp_flat_type {_ _ _} _ {t} _.
-Global Arguments mapf {_ _ _ _} _ _ {t} _.
diff --git a/src/Compilers/MapBaseType.v b/src/Compilers/MapBaseType.v
deleted file mode 100644
index 554bba4a2..000000000
--- a/src/Compilers/MapBaseType.v
+++ /dev/null
@@ -1,116 +0,0 @@
-Require Import Coq.Bool.Sumbool.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-
-Section language.
- Context {base_type_code1 base_type_code2 : Type}
- {op1 : flat_type base_type_code1 -> flat_type base_type_code1 -> Type}
- {op2 : flat_type base_type_code2 -> flat_type base_type_code2 -> Type}
- (f_base : base_type_code1 -> base_type_code2)
- (f_op : forall var s d,
- op1 s d
- -> exprf base_type_code1 op1 (var:=var) s
- -> option (op2 (lift_flat_type f_base s) (lift_flat_type f_base d))).
-
- Section with_var.
- Context {var1 : base_type_code1 -> Type}
- {var2 : base_type_code2 -> Type}
- (f_var12 : forall t, var1 t -> var2 (f_base t))
- (f_var21 : forall t, var2 (f_base t) -> var1 t)
- (failb : forall t, exprf _ op2 (var:=var2) (Tbase t)).
-
- Local Notation failf t
- := (SmartPairf (SmartValf _ failb t)).
-
- Fixpoint mapf_base_type
- {t} (e : exprf base_type_code1 op1 (var:=var1) t)
- : exprf base_type_code2 op2 (var:=var2) (lift_flat_type f_base t)
- := match e in exprf _ _ t return exprf _ _ (lift_flat_type f_base t) with
- | TT => TT
- | Var t x => Var (f_var12 _ x)
- | Op t1 tR opc args
- => let opc := f_op _ _ _ opc args in
- let args := @mapf_base_type _ args in
- match opc with
- | Some opc => Op opc args
- | None => failf _
- end
- | LetIn tx ex tC eC
- => let ex := @mapf_base_type _ ex in
- let eC := fun x => @mapf_base_type _ (eC x) in
- LetIn ex (fun x => eC (untransfer_interp_flat_type (t:=tx) f_base f_var21 x))
- | Pair tx ex ty ey
- => let ex := @mapf_base_type _ ex in
- let ey := @mapf_base_type _ ey in
- Pair ex ey
- end.
-
- Definition map_base_type
- {t} (e : expr base_type_code1 op1 t)
- : expr base_type_code2 op2 (Arrow (lift_flat_type f_base (domain t)) (lift_flat_type f_base (codomain t)))
- := let f := invert_Abs e in
- let f := fun x => mapf_base_type (f x) in
- Abs (src:=lift_flat_type f_base (domain t))
- (fun x => f (untransfer_interp_flat_type _ f_var21 x)).
- End with_var.
-
- Section bool_gen.
- Context (check_base_type : base_type_code1 -> bool)
- {var : base_type_code1 -> Type}
- (val : forall t, var t).
-
- Fixpoint check_mapf_base_type_gen
- {t} (e : exprf base_type_code1 op1 (var:=var) t)
- : bool
- := match e with
- | TT => true
- | Var t x => check_base_type t
- | Op t1 tR opc args
- => let opc := f_op _ _ _ opc args in
- let check_args := @check_mapf_base_type_gen _ args in
- match opc with
- | Some opc => check_args
- | None => false
- end
- | LetIn tx ex tC eC
- => let check_ex := @check_mapf_base_type_gen _ ex in
- let check_eC := fun x => @check_mapf_base_type_gen _ (eC x) in
- andb check_ex (check_eC (SmartValf _ val _))
- | Pair tx ex ty ey
- => let check_ex := @check_mapf_base_type_gen _ ex in
- let check_ey := @check_mapf_base_type_gen _ ey in
- andb check_ex check_ey
- end.
-
- Definition check_map_base_type_gen
- {t} (e : expr base_type_code1 op1 (var:=var) t)
- : bool
- := let f := invert_Abs e in
- let f := fun x => check_mapf_base_type_gen (f x) in
- f (SmartValf _ val _).
- End bool_gen.
-
- Section bool.
- Definition check_mapf_base_type check_base_type {t} e
- := @check_mapf_base_type_gen check_base_type (fun _ => unit) (fun _ => tt) t e.
- Definition check_map_base_type check_base_type {t} e
- := @check_map_base_type_gen check_base_type (fun _ => unit) (fun _ => tt) t e.
- End bool.
-
- Definition MapBaseType'
- (failb : forall var t, exprf _ op2 (var:=var) (Tbase t))
- {t} (e : Expr base_type_code1 op1 t)
- : Expr base_type_code2 op2 (Arrow (lift_flat_type f_base (domain t)) (lift_flat_type f_base (codomain t)))
- := fun var => map_base_type
- (var1:=fun t => var (f_base t)) (var2:=var)
- (fun _ x => x) (fun _ x => x) (failb _) (e _).
-
- Definition MapBaseType
- (failb : forall var t, exprf _ op2 (var:=var) (Tbase t))
- {t} (e : Expr base_type_code1 op1 t)
- : option (Expr base_type_code2 op2 (Arrow (lift_flat_type f_base (domain t)) (lift_flat_type f_base (codomain t))))
- := if check_map_base_type (fun _ => true (* any base type is allowed *)) (e _)
- then Some (MapBaseType' failb e)
- else None.
-End language.
diff --git a/src/Compilers/MapBaseTypeWf.v b/src/Compilers/MapBaseTypeWf.v
deleted file mode 100644
index 87c15d200..000000000
--- a/src/Compilers/MapBaseTypeWf.v
+++ /dev/null
@@ -1,117 +0,0 @@
-Require Import Coq.Bool.Sumbool.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.WfProofs.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.MapBaseType.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Prod.
-
-Section language.
- Context {base_type_code1 base_type_code2 : Type}
- {op1 : flat_type base_type_code1 -> flat_type base_type_code1 -> Type}
- {op2 : flat_type base_type_code2 -> flat_type base_type_code2 -> Type}
- (f_base : base_type_code1 -> base_type_code2)
- (f_op : forall var1 s d,
- op1 s d
- -> exprf _ op1 (var:=var1) s
- -> option (op2 (lift_flat_type f_base s) (lift_flat_type f_base d))).
-
- Local Hint Constructors wf wff or.
-
- Local Notation mapf_base_type :=
- (@mapf_base_type base_type_code1 base_type_code2 op1 op2 f_base f_op).
- Local Notation map_base_type :=
- (@map_base_type base_type_code1 base_type_code2 op1 op2 f_base f_op).
-
- Section with_var.
- Context {var1 var1' : base_type_code1 -> Type}
- {var2 var2' : base_type_code2 -> Type}
- (f_var12 : forall t, var1 t -> var2 (f_base t))
- (f_var21 : forall t, var2 (f_base t) -> var1 t)
- (f_var'12 : forall t, var1' t -> var2' (f_base t))
- (f_var'21 : forall t, var2' (f_base t) -> var1' t)
- (failb : forall t, exprf _ op2 (var:=var2) (Tbase t))
- (failb' : forall t, exprf _ op2 (var:=var2') (Tbase t))
- (Hwf_failb : forall t G, wff G (failb t) (failb' t))
- (Hwf_f_op : forall G s d opc e1 e2,
- wff G e1 e2
- -> f_op var1 s d opc e1 = f_op var1' s d opc e2)
- (Hvar12 : forall t v, f_var12 t (f_var21 t v) = v)
- (Hvar'12 : forall t v, f_var'12 t (f_var'21 t v) = v).
-
- Lemma wff_mapf_base_type G G' {t}
- (e : exprf base_type_code1 op1 (var:=var1) t)
- (e' : exprf base_type_code1 op1 (var:=var1') t)
- (HG : forall t x x',
- List.In (existT _ t (x, x')) G
- -> List.In (existT _ (f_base t) (f_var12 _ x, f_var'12 _ x')) G')
- (Hwf : wff G e e')
- : wff G'
- (mapf_base_type f_var12 f_var21 failb e)
- (mapf_base_type f_var'12 f_var'21 failb' e').
- Proof.
- revert dependent G'; induction Hwf;
- repeat first [ progress simpl in *
- | progress intros
- | progress inversion_option
- | progress subst
- | break_innermost_match_step
- | progress specialize_by_assumption
- | apply wff_SmartPairf_SmartValf
- | solve [ eauto using In_flatten_binding_list_untransfer_interp_flat_type ]
- | match goal with
- | [ |- wff _ _ _ ] => constructor
- | [ H : _ |- _ ] => apply H; try setoid_rewrite List.in_app_iff
- | [ H : f_op _ ?s ?d ?opc ?e = ?x, H' : f_op _ ?s ?d ?opc ?e' = ?y |- _ ]
- => assert (x = y) by (rewrite <- H, <- H'; eauto); clear H'
- end
- | progress destruct_head'_or ].
- Qed.
-
- Lemma wf_map_base_type {t}
- (e : expr base_type_code1 op1 (var:=var1) t)
- (e' : expr base_type_code1 op1 (var:=var1') t)
- (Hwf : wf e e')
- : wf
- (map_base_type f_var12 f_var21 failb e)
- (map_base_type f_var'12 f_var'21 failb' e').
- Proof.
- destruct Hwf; constructor; simpl; intros.
- eapply wff_mapf_base_type; [ | eauto ].
- eauto using In_flatten_binding_list_untransfer_interp_flat_type.
- Qed.
- End with_var.
-
- Section MapBaseType.
- Context (failb : forall var t, exprf _ op2 (var:=var) (Tbase t))
- (Hwf_failb : forall var1 var2 G t, wff G (failb var1 t) (failb var2 t))
- (Hwf_f_op : forall var1 var1' s d opc G e1 e2,
- wff G e1 e2
- -> f_op var1 s d opc e1 = f_op var1' s d opc e2)
- {t} (e : Expr base_type_code1 op1 t)
- (Hwf : Wf e).
-
- Lemma Wf_MapBaseType'
- : Wf (MapBaseType' f_base f_op failb e).
- Proof using Hwf Hwf_failb Hwf_f_op.
- intros var1 var2; apply wf_map_base_type; eauto.
- Qed.
-
- Lemma Wf_MapBaseType
- r
- (H : MapBaseType f_base f_op failb e = Some r)
- : Wf r.
- Proof using Hwf Hwf_failb Hwf_f_op.
- cbv [MapBaseType] in *; break_innermost_match_hyps; inversion_option; subst.
- apply Wf_MapBaseType'.
- Qed.
- End MapBaseType.
-End language.
-
-Hint Resolve @Wf_MapBaseType' : wf.
diff --git a/src/Compilers/MapCastByDeBruijn.v b/src/Compilers/MapCastByDeBruijn.v
deleted file mode 100644
index 0eb137a06..000000000
--- a/src/Compilers/MapCastByDeBruijn.v
+++ /dev/null
@@ -1,62 +0,0 @@
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.MapCast.
-Require Import Crypto.Compilers.Named.InterpretToPHOAS.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Compilers.Syntax.
-
-(** N.B. This procedure only works when there are no nested lets,
- i.e., nothing like [let x := let y := z in w] in the PHOAS syntax
- tree. This is a limitation of [compile]. *)
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t))
- {interp_base_type_bounds : base_type_code -> 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_code).
- 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))).
-
- Local Notation PContext var := (PositiveContext _ var _ base_type_code_bl_transparent).
-
- Section MapCast.
- Context {t} (e : Expr base_type_code op t)
- (input_bounds : interp_flat_type interp_base_type_bounds (domain t)).
-
- Definition MapCastCompile
- := compile (e _) (DefaultNamesFor e).
- Definition MapCastDoCast (e' : option (Named.expr base_type_code op BinNums.positive t))
- := option_map
- (fun e'' => map_cast
- interp_op_bounds pick_typeb cast_op
- (BoundsContext:=PContext _)
- empty
- e''
- input_bounds)
- e'.
- Definition MapCastDoInterp
- (e' : option
- (option
- { output_bounds : interp_flat_type interp_base_type_bounds (codomain t) &
- Named.expr _ _ _ (Arrow (pick_type input_bounds) (pick_type output_bounds)) }))
- : option { output_bounds : interp_flat_type interp_base_type_bounds (codomain t)
- & Expr base_type_code op (Arrow (pick_type input_bounds) (pick_type output_bounds)) }
- := match e' with
- | Some (Some (existT output_bounds e''))
- => Some (existT _ output_bounds (InterpToPHOAS (Context:=fun var => PContext var) failb e''))
- | Some None | None => None
- end.
- Definition MapCast
- : option { output_bounds : interp_flat_type interp_base_type_bounds (codomain t)
- & Expr base_type_code op (Arrow (pick_type input_bounds) (pick_type output_bounds)) }
- := MapCastDoInterp (MapCastDoCast MapCastCompile).
- End MapCast.
-End language.
diff --git a/src/Compilers/MapCastByDeBruijnInterp.v b/src/Compilers/MapCastByDeBruijnInterp.v
deleted file mode 100644
index 9e7454c2a..000000000
--- a/src/Compilers/MapCastByDeBruijnInterp.v
+++ /dev/null
@@ -1,129 +0,0 @@
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.MapCastInterp.
-Require Import Crypto.Compilers.Named.MapCastWf.
-Require Import Crypto.Compilers.Named.InterpretToPHOASInterp.
-Require Import Crypto.Compilers.Named.CompileInterp.
-Require Import Crypto.Compilers.Named.CompileInterpSideConditions.
-Require Import Crypto.Compilers.Named.CompileWf.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Compilers.Named.PositiveContext.DefaultsProperties.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.MapCastByDeBruijn.
-Require Import Crypto.Compilers.InterpSideConditions.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t))
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst)
- {interp_base_type_bounds : base_type_code -> 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)
- (interped_op_side_conditions : forall s d, op s d -> interp_flat_type interp_base_type s -> pointed_Prop)
- (pick_typeb : forall t, interp_base_type_bounds t -> base_type_code).
- 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)
- (Hside : to_prop (interped_op_side_conditions _ _ opc 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))
- (Hside : to_prop (interped_op_side_conditions _ _ opc (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
- base_type_code op base_type_code_beq base_type_code_bl_transparent
- failb interp_base_type_bounds interp_op_bounds pick_typeb cast_op).
-
- Local Notation PositiveContextOk := (@PositiveContextOk base_type_code _ base_type_code_beq base_type_code_bl_transparent base_type_code_lb).
-
- Local Instance dec_base_type_code_eq : DecidableRel (@eq base_type_code).
- Proof.
- refine (fun x y => (if base_type_code_beq x y as b return base_type_code_beq x y = b -> Decidable (x = y)
- then fun pf => left (base_type_code_bl_transparent _ _ pf)
- else fun pf => right _) eq_refl).
- { clear -pf base_type_code_lb.
- let pf := pf in
- abstract (intro; erewrite base_type_code_lb in pf by eassumption; congruence). }
- Defined.
-
- Local Arguments Compile.compile : simpl never.
- Lemma MapCastCorrect
- {t} (e : Expr base_type_code 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)
- (Hside : to_prop (InterpSideConditions interp_op interped_op_side_conditions e 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 base_type_code_lb interp_op_bounds_correct pull_cast_back.
- unfold MapCastByDeBruijn.MapCast, MapCastCompile, MapCastDoCast, MapCastDoInterp, option_map; intros b e'.
- break_innermost_match; try congruence; intros ? v v'.
- inversion_option; inversion_sigma; subst; simpl in *; intros.
- lazymatch goal with
- | [ H : MapCast.map_cast _ _ _ _ _ _ = Some _ |- _ ]
- => eapply map_cast_correct with (t:=Arrow _ _) (oldValues:=empty) (newValues:=empty) in H;
- [ destruct H; split; [ | eassumption ] | try eassumption.. ]
- end;
- try solve [ eassumption
- | auto using PositiveContextOk with typeclass_instances
- | repeat first [ rewrite !lookupb_empty by (apply PositiveContextOk; assumption)
- | intro
- | congruence ] ];
- unfold Interp;
- [ match goal with
- | [ H : ?y = Some ?b |- ?x = ?b ]
- => cut (y = Some x); [ congruence | ]
- end
- |
- | change (interp interp_op (?e ?var) ?v') with (Interp interp_op e v');
- unfold Interp, InterpretToPHOAS.Named.InterpToPHOAS, InterpretToPHOAS.Named.InterpToPHOAS_gen;
- rewrite <- interp_interp_to_phoas; [ reflexivity | ]
- | ].
- { erewrite (interp_compile (ContextOk:=PositiveContextOk)) with (e':=e _);
- [ reflexivity | auto | .. | eassumption ];
- auto using name_list_unique_DefaultNamesFor. }
- { erewrite (interp_compile (ContextOk:=PositiveContextOk)) with (e':=e _);
- [ reflexivity | auto | .. | eassumption ];
- auto using name_list_unique_DefaultNamesFor. }
- { intro; eapply wf_map_cast with (t := Arrow _ _) (fValues := empty); eauto using PositiveContextOk with typeclass_instances.
- { eapply (wf_compile (ContextOk:=PositiveContextOk) (make_var':=fun _ => id)) with (e':= e _);
- [ auto | .. | eassumption ];
- auto using name_list_unique_DefaultNamesFor. }
- { intros ???; rewrite lookupb_empty by apply PositiveContextOk; congruence. } }
- { erewrite (interp_side_conditions_compile (ContextOk:=PositiveContextOk)) with (e':=e _);
- [ assumption | auto | .. | eassumption ];
- auto using name_list_unique_DefaultNamesFor. }
- Qed.
-End language.
diff --git a/src/Compilers/MapCastByDeBruijnWf.v b/src/Compilers/MapCastByDeBruijnWf.v
deleted file mode 100644
index d611f50b6..000000000
--- a/src/Compilers/MapCastByDeBruijnWf.v
+++ /dev/null
@@ -1,108 +0,0 @@
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.MapCastWf.
-Require Import Crypto.Compilers.Named.InterpretToPHOASWf.
-Require Import Crypto.Compilers.Named.CompileWf.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Compilers.Named.PositiveContext.DefaultsProperties.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.MapCastByDeBruijn.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true)
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t))
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst)
- {interp_base_type_bounds : base_type_code -> 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_code).
- 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
- base_type_code op base_type_code_beq base_type_code_bl_transparent
- failb interp_base_type_bounds interp_op_bounds pick_typeb cast_op).
-
- Local Notation PositiveContextOk := (@PositiveContextOk base_type_code _ base_type_code_beq base_type_code_bl_transparent base_type_code_lb).
-
- Local Instance dec_base_type_code_eq : DecidableRel (@eq base_type_code).
- Proof.
- refine (fun x y => (if base_type_code_beq x y as b return base_type_code_beq x y = b -> Decidable (x = y)
- then fun pf => left (base_type_code_bl_transparent _ _ pf)
- else fun pf => right _) eq_refl).
- { clear -pf base_type_code_lb.
- let pf := pf in
- abstract (intro; erewrite base_type_code_lb in pf by eassumption; congruence). }
- Defined.
-
- Local Arguments Compile.compile : simpl never.
- Lemma Wf_MapCast
- {t} (e : Expr base_type_code op t)
- (input_bounds : interp_flat_type interp_base_type_bounds (domain t))
- : forall {b} e' (He':MapCast e input_bounds = Some (existT _ b e')) (Hwf : Wf e),
- Wf e'.
- Proof using base_type_code_lb.
- unfold MapCastByDeBruijn.MapCast, MapCastCompile, MapCastDoCast, MapCastDoInterp, option_map; intros b e'.
- break_innermost_match; try congruence; intros ? v v'.
- inversion_option; inversion_sigma; subst; simpl in *; intros.
- unfold InterpretToPHOAS.Named.InterpToPHOAS, InterpretToPHOAS.Named.InterpToPHOAS_gen.
- destruct t as [src dst].
- eapply (@wf_interp_to_phoas
- base_type_code op FMapPositive.PositiveMap.key _ _ _ _
- (PositiveContext base_type_code _ base_type_code_beq base_type_code_bl_transparent)
- (PositiveContext base_type_code _ base_type_code_beq base_type_code_bl_transparent)
- PositiveContextOk PositiveContextOk
- (failb _) (failb _) _ e1);
- (eapply wf_map_cast with (fValues:=empty); eauto using PositiveContextOk with typeclass_instances);
- try (eapply (wf_compile (make_var':=fun _ => id) (ContextOk:=PositiveContextOk));
- [ eauto
- | ..
- | eassumption ]);
- try solve [ auto using name_list_unique_DefaultNamesFor
- | intros ???; rewrite lookupb_empty by apply PositiveContextOk; congruence ].
- Qed.
-
- Lemma Wf_MapCast_arrow
- {s d} (e : Expr base_type_code op (Arrow s d))
- (input_bounds : interp_flat_type interp_base_type_bounds s)
- : forall {b} (e' : Expr _ _ (Arrow (pick_type input_bounds) (pick_type b)))
- (He':MapCast e input_bounds = Some (existT _ b e'))
- (Hwf : Wf e),
- Wf e'.
- Proof using base_type_code_lb. exact (@Wf_MapCast (Arrow s d) e input_bounds). Qed.
-End language.
-
-Hint Resolve Wf_MapCast Wf_MapCast_arrow : wf.
diff --git a/src/Compilers/MultiSizeTest.v b/src/Compilers/MultiSizeTest.v
deleted file mode 100644
index c31127126..000000000
--- a/src/Compilers/MultiSizeTest.v
+++ /dev/null
@@ -1,277 +0,0 @@
-Require Import Coq.omega.Omega.
-Require Import Crypto.Compilers.SmartMap.
-
-Set Implicit Arguments.
-Set Asymmetric Patterns.
-
-(** * Preliminaries: bounded and unbounded number types *)
-
-Definition bound8 := 256.
-
-Definition word8 := {n | n < bound8}.
-
-Definition bound9 := 512.
-
-Definition word9 := {n | n < bound9}.
-
-
-(** * Expressions over unbounded words *)
-
-Section unbounded.
- Variable var : Type.
-
- Inductive unbounded :=
- | Const : nat -> unbounded
- | Var : var -> unbounded
- | Plus : unbounded -> unbounded -> unbounded
- | LetIn : unbounded -> (var -> unbounded) -> unbounded.
-End unbounded.
-
-Arguments Const [var] _.
-
-Definition Unbounded := forall var, unbounded var.
-
-Fixpoint unboundedD (e : unbounded nat) : nat :=
- match e with
- | Const n => n
- | Var n => n
- | Plus e1 e2 => unboundedD e1 + unboundedD e2
- | LetIn e1 e2 => unboundedD (e2 (unboundedD e1))
- end.
-
-Definition UnboundedD (E : Unbounded) : nat :=
- unboundedD (E _).
-
-(** * Opt-in bounded types *)
-
-Section bounded.
- Inductive type :=
- | Nat
- | Word8
- | Word9.
-
- Variable var : type -> Type.
-
- Inductive bounded : type -> Type :=
- | BConst : nat -> bounded Nat
- | BConst8 : word8 -> bounded Word8
- | BConst9 : word9 -> bounded Word9
- | BVar : forall t, var t -> bounded t
- | BPlus : bounded Nat -> bounded Nat -> bounded Nat
- | BPlus8 : bounded Word8 -> bounded Word8 -> bounded Word8
- | BPlus9 : bounded Word9 -> bounded Word9 -> bounded Word9
- | BLetIn : forall t1 t2, bounded t1 -> (var t1 -> bounded t2) -> bounded t2
-
- | Unbound : forall t, bounded t -> bounded Nat
- | Bound : forall t, bounded Nat -> bounded t.
-End bounded.
-
-Arguments BConst [var] _.
-Arguments BConst8 [var] _.
-Arguments BConst9 [var] _.
-Arguments BVar [var t] _.
-Arguments Unbound [var t] _.
-Arguments Bound [var] _ _.
-
-Definition Bounded t := forall var, bounded var t.
-
-Definition typeD (t : type) : Type :=
- match t with
- | Nat => nat
- | Word8 => word8
- | Word9 => word9
- end.
-
-Theorem O_lt_S : forall n, O < S n.
-Proof.
- intros; omega.
-Qed.
-
-Definition plus8 (a b : word8) : word8 :=
- let n := proj1_sig a + proj1_sig b in
- match le_lt_dec bound8 n with
- | left _ => exist _ O (O_lt_S _)
- | right pf => exist _ n pf
- end.
-
-Definition plus9 (a b : word9) : word9 :=
- let n := proj1_sig a + proj1_sig b in
- match le_lt_dec bound9 n with
- | left _ => exist _ O (O_lt_S _)
- | right pf => exist _ n pf
- end.
-
-Infix "+8" := plus8 (at level 50).
-Infix "+9" := plus9 (at level 50).
-
-Definition unbound {t} : typeD t -> nat :=
- match t with
- | Nat => fun x => x
- | Word8 => fun x => proj1_sig x
- | Word9 => fun x => proj1_sig x
- end.
-
-Definition bound {t} : nat -> typeD t :=
- match t return nat -> typeD t with
- | Nat => fun x => x
- | Word8 => fun x =>
- match le_lt_dec bound8 x with
- | left _ => exist _ O (O_lt_S _)
- | right pf => exist _ x pf
- end
- | Word9 => fun x =>
- match le_lt_dec bound9 x with
- | left _ => exist _ O (O_lt_S _)
- | right pf => exist _ x pf
- end
- end.
-
-Fixpoint boundedD t (e : bounded typeD t) : typeD t :=
- match e with
- | BConst n => n
- | BConst8 n => n
- | BConst9 n => n
- | BVar _ n => n
- | BPlus e1 e2 => boundedD e1 + boundedD e2
- | BPlus8 e1 e2 => boundedD e1 +8 boundedD e2
- | BPlus9 e1 e2 => boundedD e1 +9 boundedD e2
- | BLetIn _ _ e1 e2 => boundedD (e2 (boundedD e1))
- | Unbound _ e1 => unbound (boundedD e1)
- | Bound _ e1 => bound (boundedD e1)
- end.
-
-Definition BoundedD t (E : Bounded t) : typeD t :=
- boundedD (E _).
-
-
-(** * Insertion of bounded types opportunistically *)
-
-Definition fail {var} : nat * bounded var Nat := (0, BConst 0).
-
-Fixpoint boundOf (eb : unbounded nat) : nat :=
- match eb with
- | Const n => n
- | Var n => n
- | Plus eb1 eb2 => boundOf eb1 + boundOf eb2
- | LetIn eb1 eb2 => boundOf (eb2 (boundOf eb1))
- end.
-
-Fixpoint boundify {var} (eb : unbounded nat) (e : unbounded (var Nat)) : nat * bounded var Nat :=
- match e with
- | Const n => (n,
- match le_lt_dec bound8 n with
- | left _ =>
- match le_lt_dec bound9 n with
- | left _ => BConst n
- | right pf => Unbound (BConst9 (exist _ n pf))
- end
- | right pf => Unbound (BConst8 (exist _ n pf))
- end)
- | Var x =>
- match eb with
- | Var n => (n, BVar x)
- | _ => fail
- end
- | Plus e1 e2 =>
- match eb with
- | Plus eb1 eb2 =>
- let (n1, e1') := boundify eb1 e1 in
- let (n2, e2') := boundify eb2 e2 in
- (n1 + n2,
- if le_lt_dec bound8 (n1 + n2)
- then if le_lt_dec bound9 (n1 + n2)
- then BPlus e1' e2'
- else Unbound (BPlus9 (Bound _ e1') (Bound _ e2'))
- else Unbound (BPlus8 (Bound _ e1') (Bound _ e2')))
- | _ => fail
- end
- | LetIn e1 e2 =>
- match eb with
- | LetIn eb1 eb2 =>
- let (n1, e1') := boundify eb1 e1 in
- (boundOf (eb2 n1), BLetIn e1' (fun x => snd (boundify (eb2 n1) (e2 x))))
- | _ => fail
- end
- end.
-
-Definition Boundify (E : Unbounded) : Bounded Nat :=
- fun _ => snd (boundify (E _) (E _)).
-
-
-(** * Moving [Unbound] operators down from [LetIn]s to their use sites *)
-
-Fixpoint movedown {var t} (e : bounded (bounded var) t) : bounded var t :=
- match e with
- | BConst n => BConst n
- | BConst8 n => BConst8 n
- | BConst9 n => BConst9 n
- | BVar _ e => e
- | BPlus e1 e2 => BPlus (movedown e1) (movedown e2)
- | BPlus8 e1 e2 => BPlus8 (movedown e1) (movedown e2)
- | BPlus9 e1 e2 => BPlus9 (movedown e1) (movedown e2)
- | BLetIn _ _ e1 e2 =>
- match movedown e1 in bounded _ t return (bounded _ t -> _) -> _ with
- | Unbound _ e1'' => fun e2_rec => BLetIn e1'' (fun x => e2_rec (Unbound (BVar x)))
- | e1' => fun e2_rec => BLetIn e1' (fun x => e2_rec (BVar x))
- end (fun x => movedown (e2 x))
- | Unbound _ e1 => Unbound (movedown e1)
- | Bound t e1 => Bound t (movedown e1)
- end.
-
-Definition Movedown t (E : Bounded t) : Bounded t :=
- fun _ => movedown (E _).
-
-
-(** * Canceling matching [Bound] and [Unbound] *)
-
-Definition type_eq_dec : forall t1 t2 : type, {t1 = t2} + {t1 <> t2}.
-Proof.
- decide equality.
-Defined.
-
-Fixpoint cancel {var t} (e : bounded var t) : bounded var t :=
- match e with
- | BConst n => BConst n
- | BConst8 n => BConst8 n
- | BConst9 n => BConst9 n
- | BVar _ x => BVar x
- | BPlus e1 e2 => BPlus (cancel e1) (cancel e2)
- | BPlus8 e1 e2 => BPlus8 (cancel e1) (cancel e2)
- | BPlus9 e1 e2 => BPlus9 (cancel e1) (cancel e2)
- | BLetIn _ _ e1 e2 => BLetIn (cancel e1) (fun x => cancel (e2 x))
- | Unbound _ e1 => Unbound (cancel e1)
- | Bound t e1 =>
- match cancel e1 with
- | Unbound t' e1' =>
- match type_eq_dec t' t with
- | left pf => match pf in _ = T return bounded _ T with
- | eq_refl => e1'
- end
- | right _ => Bound t (Unbound e1')
- end
- | e1' => Bound t e1'
- end
- end.
-
-Definition Cancel t (E : Bounded t) : Bounded t :=
- fun _ => cancel (E _).
-
-
-(** * Examples *)
-
-Example ex1 : Unbounded := fun _ =>
- LetIn (Const 127) (fun a =>
- LetIn (Const 63) (fun b =>
- LetIn (Plus (Var a) (Var b)) (fun c =>
- Plus (Var c) (Var c)))).
-
-Eval compute in (UnboundedD ex1).
-
-Definition ex1b := Boundify ex1.
-Eval compute in ex1b.
-
-Definition ex1bm := Movedown (Boundify ex1).
-Eval compute in ex1bm.
-
-Definition ex1bmc := Cancel (Movedown (Boundify ex1)).
-Eval compute in ex1bmc.
diff --git a/src/Compilers/Named/AListContext.v b/src/Compilers/Named/AListContext.v
deleted file mode 100644
index ffc8355ee..000000000
--- a/src/Compilers/Named/AListContext.v
+++ /dev/null
@@ -1,143 +0,0 @@
-(** * Context made from an associative list, without modules *)
-Require Import Coq.Bool.Sumbool.
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Equality.
-
-Local Open Scope list_scope.
-Section ctx.
- Context (key : Type)
- (key_beq : key -> key -> bool)
- (key_bl : forall k1 k2, key_beq k1 k2 = true -> k1 = k2)
- (key_lb : forall k1 k2, k1 = k2 -> key_beq k1 k2 = true)
- base_type_code (var : base_type_code -> Type)
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true).
-
- Definition var_cast {a b} (x : var a) : option (var b)
- := match Sumbool.sumbool_of_bool (base_type_code_beq a b), Sumbool.sumbool_of_bool (base_type_code_beq b b) with
- | left pf, left pf' => match eq_trans (base_type_code_bl_transparent _ _ pf) (eq_sym (base_type_code_bl_transparent _ _ pf')) with
- | eq_refl => Some x
- end
- | right _, _ | _, right _ => None
- end.
-
- Fixpoint find (k : key) (xs : list (key * { t : _ & var t })) {struct xs}
- : option { t : _ & var t }
- := match xs with
- | nil => None
- | k'x :: xs' =>
- if key_beq k (fst k'x)
- then Some (snd k'x)
- else find k xs'
- end.
-
- Fixpoint remove (k : key) (xs : list (key * { t : _ & var t })) {struct xs}
- : list (key * { t : _ & var t })
- := match xs with
- | nil => nil
- | k'x :: xs' =>
- if key_beq k (fst k'x)
- then remove k xs'
- else k'x :: remove k xs'
- end.
-
- Definition add (k : key) (x : { t : _ & var t }) (xs : list (key * { t : _ & var t }))
- : list (key * { t : _ & var t })
- := (k, x) :: xs.
-
- Lemma find_remove_neq k k' xs (H : k <> k')
- : find k (remove k' xs) = find k xs.
- Proof.
- induction xs as [|x xs IHxs]; [ reflexivity | simpl ].
- break_innermost_match;
- repeat match goal with
- | [ H : key_beq _ _ = true |- _ ] => apply key_bl in H
- | [ H : context[key_beq ?x ?x] |- _ ] => rewrite (key_lb x x) in H by reflexivity
- | [ |- context[key_beq ?x ?x] ] => rewrite (key_lb x x) by reflexivity
- | [ H : ?x = false |- context[?x] ] => rewrite H
- | _ => congruence
- | _ => assumption
- | _ => progress subst
- | _ => progress simpl
- end.
- Qed.
-
- Lemma find_remove_same k xs
- : find k (remove k xs) = None.
- Proof.
- induction xs as [|x xs IHxs]; [ reflexivity | simpl ].
- break_innermost_match;
- repeat match goal with
- | [ H : key_beq _ _ = true |- _ ] => apply key_bl in H
- | [ H : context[key_beq ?x ?x] |- _ ] => rewrite (key_lb x x) in H by reflexivity
- | [ |- context[key_beq ?x ?x] ] => rewrite (key_lb x x) by reflexivity
- | [ H : ?x = false |- context[?x] ] => rewrite H
- | _ => congruence
- | _ => assumption
- | _ => progress subst
- | _ => progress simpl
- end.
- Qed.
-
- Lemma find_remove_nbeq k k' xs (H : key_beq k k' = false)
- : find k (remove k' xs) = find k xs.
- Proof.
- rewrite find_remove_neq; [ reflexivity | intro; subst ].
- rewrite key_lb in H by reflexivity; congruence.
- Qed.
-
- Lemma find_remove_beq k k' xs (H : key_beq k k' = true)
- : find k (remove k' xs) = None.
- Proof.
- apply key_bl in H; subst.
- rewrite find_remove_same; reflexivity.
- Qed.
-
- Definition AListContext : @Context base_type_code key var
- := {| ContextT := list (key * { t : _ & var t });
- lookupb t ctx n
- := match find n ctx with
- | Some (existT t' v)
- => var_cast v
- | None => None
- end;
- extendb t ctx n v
- := add n (existT _ t v) ctx;
- removeb t ctx n
- := remove n ctx;
- empty := nil |}.
-
- Lemma length_extendb (ctx : AListContext) k t v
- : length (@extendb _ _ _ AListContext t ctx k v) = S (length ctx).
- Proof. reflexivity. Qed.
-
- Lemma AListContextOk : @ContextOk base_type_code key var AListContext.
- Proof using base_type_code_lb key_bl key_lb.
- split;
- repeat first [ reflexivity
- | progress simpl in *
- | progress intros
- | rewrite find_remove_nbeq by eassumption
- | rewrite find_remove_beq by eassumption
- | rewrite find_remove_neq by congruence
- | rewrite find_remove_same by congruence
- | match goal with
- | [ |- context[key_beq ?x ?y] ]
- => destruct (key_beq x y) eqn:?
- | [ H : key_beq ?x ?y = true |- _ ]
- => apply key_bl in H
- end
- | break_innermost_match_step
- | progress unfold var_cast
- | rewrite key_lb in * by reflexivity
- | rewrite base_type_code_lb in * by reflexivity
- | rewrite concat_pV
- | congruence
- | break_innermost_match_hyps_step
- | progress unfold var_cast in * ].
- Qed.
-End ctx.
diff --git a/src/Compilers/Named/Compile.v b/src/Compilers/Named/Compile.v
deleted file mode 100644
index bee71cea5..000000000
--- a/src/Compilers/Named/Compile.v
+++ /dev/null
@@ -1,59 +0,0 @@
-(** * PHOAS → Named Representation of Gallina *)
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.Syntax.
-
-Local Notation eta x := (fst x, snd x).
-
-Local Open Scope ctype_scope.
-Local Open Scope nexpr_scope.
-Local Open Scope expr_scope.
-Section language.
- Context (base_type_code : Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (Name : Type).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op (fun _ => Name)).
- Local Notation expr := (@expr base_type_code op (fun _ => Name)).
- Local Notation nexprf := (@Named.exprf base_type_code op Name).
- Local Notation nexpr := (@Named.expr base_type_code op Name).
-
- Fixpoint ocompilef {t} (e : exprf t) (ls : list (option Name)) {struct e}
- : option (nexprf t)
- := match e in @Syntax.exprf _ _ _ t return option (nexprf t) with
- | TT => Some Named.TT
- | Var _ x => Some (Named.Var x)
- | Op _ _ op args => option_map (Named.Op op) (@ocompilef _ args ls)
- | LetIn tx ex _ eC
- => match @ocompilef _ ex nil, split_onames tx ls with
- | Some x, (Some n, ls')%core
- => option_map (fun C => Named.LetIn tx n x C) (@ocompilef _ (eC n) ls')
- | _, _ => None
- end
- | Pair _ ex _ ey => match @ocompilef _ ex nil, @ocompilef _ ey nil with
- | Some x, Some y => Some (Named.Pair x y)
- | _, _ => None
- end
- end.
-
- Definition ocompile {t} (e : expr t) (ls : list (option Name))
- : option (nexpr t)
- := match e in @Syntax.expr _ _ _ t return option (nexpr t) with
- | Abs src _ f
- => match split_onames src ls with
- | (Some n, ls')%core
- => option_map (Named.Abs n) (@ocompilef _ (f n) ls')
- | _ => None
- end
- end.
-
- Definition compilef {t} (e : exprf t) (ls : list Name) := @ocompilef t e (List.map (@Some _) ls).
- Definition compile {t} (e : expr t) (ls : list Name) := @ocompile t e (List.map (@Some _) ls).
-End language.
-
-Global Arguments ocompilef {_ _ _ _} e ls.
-Global Arguments ocompile {_ _ _ _} e ls.
-Global Arguments compilef {_ _ _ _} e ls.
-Global Arguments compile {_ _ _ _} e ls.
diff --git a/src/Compilers/Named/CompileInterp.v b/src/Compilers/Named/CompileInterp.v
deleted file mode 100644
index 20a536ddc..000000000
--- a/src/Compilers/Named/CompileInterp.v
+++ /dev/null
@@ -1,207 +0,0 @@
-(** * PHOAS → Named Representation of Gallina *)
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.Named.NameUtilProperties.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.NameUtil.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-
-Local Open Scope ctype_scope.
-Local Open Scope nexpr_scope.
-Local Open Scope expr_scope.
-Section language.
- Context {base_type_code}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {interp_base_type : base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}
- {base_type_dec : DecidableRel (@eq base_type_code)}
- {Name_dec : DecidableRel (@eq Name)}
- {Context : @Context base_type_code Name interp_base_type}
- {ContextOk : ContextOk Context}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op (fun _ => Name)).
- Local Notation expr := (@expr base_type_code op (fun _ => Name)).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op (fun _ => Name) interp_base_type).
- Local Notation wf := (@wf base_type_code op (fun _ => Name) interp_base_type).
- Local Notation nexprf := (@Named.exprf base_type_code op Name).
- Local Notation nexpr := (@Named.expr base_type_code op Name).
- Local Notation ocompilef := (@ocompilef base_type_code op Name).
- Local Notation ocompile := (@ocompile base_type_code op Name).
- Local Notation compilef := (@compilef base_type_code op Name).
- Local Notation compile := (@compile base_type_code op Name).
-
- Lemma interpf_ocompilef (ctx : Context) {t} (e : exprf t) e' (ls : list (option Name))
- G
- (Hwf : wff G e e')
- (HG : forall t n x, List.In (existT _ t (n, x)%core) G -> lookupb t ctx n = Some x)
- v
- (H : ocompilef e ls = Some v)
- (Hls : oname_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In (Some n) ls -> False)
- : Named.interpf (interp_op:=interp_op) (ctx:=ctx) v
- = Some (interpf interp_op e').
- Proof using ContextOk Name_dec base_type_dec.
- revert dependent ctx; revert dependent ls; induction Hwf;
- repeat first [ progress intros
- | progress subst
- | progress inversion_option
- | apply (f_equal (@Some _))
- | apply (f_equal (@interp_op _ _ _))
- | solve [ eauto ]
- | progress simpl in *
- | progress unfold option_map, LetIn.Let_In in *
- | progress break_innermost_match_step
- | progress break_match_hyps
- | progress destruct_head' or
- | progress inversion_prod
- | progress specialize_by_assumption
- | progress specialize_by auto using oname_list_unique_nil
- | match goal with
- | [ H : forall x, oname_list_unique ?ls -> _ |- _ ]
- => specialize (fun pf x => H x pf)
- | [ H : context[snd (split_onames _ _)] |- _ ]
- => rewrite snd_split_onames_skipn in H
- | [ H : oname_list_unique (List.skipn _ _) -> _ |- _ ]
- => specialize (fun pf => H (@oname_list_unique_skipn _ _ _ pf))
- | [ IH : forall v ls, ocompilef ?e ls = Some v -> _, H : ocompilef ?e ?ls' = Some ?v' |- _ ]
- => specialize (IH _ _ H)
- | [ IH : forall x1 x2 v ls, ocompilef (?e x1) ls = Some v -> _, H : ocompilef (?e ?x1') ?ls' = Some ?v' |- _ ]
- => specialize (fun x2 => IH _ x2 _ _ H)
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => rewrite List.in_app_iff in H
- | [ H : forall ctx, _ -> Named.interpf ?e = Some _, H' : context[Named.interpf ?e] |- _ ]
- => rewrite H in H' by assumption
- | [ H : forall x2 ctx, _ -> Named.interpf ?e = Some _ |- Named.interpf ?e = Some _ ]
- => apply H; clear H
- | [ H : forall x2, _ -> forall ctx, _ -> Named.interpf ?e = Some _ |- Named.interpf ?e = Some _ ]
- => apply H; clear H
- end ];
- repeat match goal with
- | _ => erewrite lookupb_extend by assumption
- | [ |- context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- | [ H : _ |- _ ] => erewrite H by eassumption
- | _ => progress unfold dec in *
- | _ => progress break_innermost_match_step
- | _ => progress subst
- | _ => progress destruct_head' and
- | _ => congruence
- | [ H : List.In _ (flatten_binding_list _ _) |- _ ]
- => erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) in H;
- [ | | apply path_prod_uncurried; split; [ eassumption | simpl; reflexivity ] ];
- [ | solve [ eauto using oname_list_unique_firstn, oname_list_unique_skipn ] ]
- | [ H : _ |- _ ]
- => first [ erewrite find_Name_and_val_wrong_type in H by eassumption
- | rewrite find_Name_and_val_different in H by assumption
- | rewrite snd_split_onames_skipn in H ]
- | _ => solve [ eauto using In_skipn, In_firstn
- | eapply split_onames_find_Name_Some_unique; [ | apply path_prod; simpl | ]; eauto ]
- | [ H : find_Name_and_val _ _ ?t ?n ?N ?V None = Some _, H' : List.In (Some ?n) (List.skipn _ ?ls) |- False ]
- => eapply find_Name_and_val_find_Name_Some, split_onames_find_Name_Some_unique in H;
- [ | | apply path_prod_uncurried; split; [ eassumption | simpl; reflexivity ] ];
- [ | solve [ eauto using oname_list_unique_firstn, oname_list_unique_skipn ] ]
- | [ H : List.In (existT _ ?t (?n, _)%core) ?G,
- H' : List.In (Some ?n) (List.skipn _ ?ls),
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> List.In (Some n') ?ls -> False
- |- False ]
- => apply (IH _ _ _ H); clear IH H
- | [ H : List.In (existT _ ?t (?n, _)%core) ?G,
- H' : find_Name _ ?n ?N = Some ?t',
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> List.In (Some n') ?ls -> False
- |- _ ]
- => exfalso; apply (IH _ _ _ H); clear IH H
- end.
- Qed.
-
- Lemma interp_ocompile (ctx : Context) {t} (e : expr t) e' (ls : list (option Name))
- (Hwf : wf e e')
- f
- (Hls : oname_list_unique ls)
- (H : ocompile e ls = Some f)
- : forall v, Named.interp (interp_op:=interp_op) (ctx:=ctx) f v
- = Some (interp interp_op e' v).
- Proof using ContextOk Name_dec base_type_dec.
- revert H; destruct Hwf;
- repeat first [ progress simpl in *
- | progress unfold option_map, Named.interp in *
- | congruence
- | progress break_innermost_match
- | progress inversion_option
- | progress subst
- | progress intros ].
- eapply interpf_ocompilef;
- [ eauto | | eassumption
- | inversion_prod; subst; rewrite snd_split_onames_skipn; eauto using oname_list_unique_skipn
- |intros ???; erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) by eassumption;
- let H := fresh in
- intro H; apply find_Name_and_val_find_Name_Some in H;
- eapply split_onames_find_Name_Some_unique in H; [ | eassumption.. ];
- intuition ].
- { intros ???.
- repeat first [ solve [ auto ]
- | rewrite (lookupb_extend _ _ _)
- | progress subst
- | progress break_innermost_match
- | erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) by eassumption
- | congruence
- | match goal with
- | [ |- context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- | [ H : _ |- _ ] => first [ erewrite find_Name_and_val_wrong_type in H by eassumption
- | erewrite find_Name_and_val_different in H by eassumption ]
- end
- | progress intros ]. }
- Qed.
-
- Lemma interpf_compilef (ctx : Context) {t} (e : exprf t) e' (ls : list Name)
- G
- (Hwf : wff G e e')
- (HG : forall t n x, List.In (existT _ t (n, x)%core) G -> lookupb t ctx n = Some x)
- v
- (H : compilef e ls = Some v)
- (Hls : name_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In n ls -> False)
- : Named.interpf (interp_op:=interp_op) (ctx:=ctx) v
- = Some (interpf interp_op e').
- Proof using ContextOk Name_dec base_type_dec.
- eapply interpf_ocompilef; try eassumption.
- setoid_rewrite List.in_map_iff; intros; destruct_head' ex; destruct_head' and; inversion_option; subst.
- eauto.
- Qed.
-
- Lemma interp_compile (ctx : Context) {t} (e : expr t) e' (ls : list Name)
- (Hwf : wf e e')
- f
- (Hls : name_list_unique ls)
- (H : compile e ls = Some f)
- : forall v, Named.interp (interp_op:=interp_op) (ctx:=ctx) f v
- = Some (interp interp_op e' v).
- Proof using ContextOk Name_dec base_type_dec. eapply interp_ocompile; eassumption. Qed.
-
- Lemma Interp_compile {t} (e : Expr t) (ls : list Name)
- (Hwf : Wf e)
- f
- (Hls : name_list_unique ls)
- (H : compile (e _) ls = Some f)
- : forall v, Named.Interp (Context:=Context) (interp_op:=interp_op) f v
- = Some (Interp interp_op e v).
- Proof using ContextOk Name_dec base_type_dec. eapply interp_compile; eauto. Qed.
-End language.
diff --git a/src/Compilers/Named/CompileInterpSideConditions.v b/src/Compilers/Named/CompileInterpSideConditions.v
deleted file mode 100644
index da8e1d8b2..000000000
--- a/src/Compilers/Named/CompileInterpSideConditions.v
+++ /dev/null
@@ -1,253 +0,0 @@
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.Named.NameUtilProperties.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.NameUtil.
-Require Import Crypto.Compilers.Named.InterpSideConditions.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.InterpSideConditions.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-
-Section language.
- Context {base_type_code}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {interp_base_type : base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}
- {base_type_dec : DecidableRel (@eq base_type_code)}
- {Name_dec : DecidableRel (@eq Name)}
- {interped_op_side_conditions : forall s d, op s d -> interp_flat_type interp_base_type s -> pointed_Prop}
- {Context : @Context base_type_code Name interp_base_type}
- {ContextOk : ContextOk Context}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op (fun _ => Name)).
- Local Notation expr := (@expr base_type_code op (fun _ => Name)).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op (fun _ => Name) interp_base_type).
- Local Notation wf := (@wf base_type_code op (fun _ => Name) interp_base_type).
- Local Notation nexprf := (@Named.exprf base_type_code op Name).
- Local Notation nexpr := (@Named.expr base_type_code op Name).
- Local Notation ocompilef := (@ocompilef base_type_code op Name).
- Local Notation ocompile := (@ocompile base_type_code op Name).
- Local Notation compilef := (@compilef base_type_code op Name).
- Local Notation compile := (@compile base_type_code op Name).
-
- Lemma interpf_side_conditions_gen_ocompilef (ctx : Context) {t} (e : exprf t) e' (ls : list (option Name))
- G
- (Hwf : wff G e e')
- (HG : forall t n x, List.In (existT _ t (n, x)%core) G -> lookupb t ctx n = Some x)
- v
- (H : ocompilef e ls = Some v)
- (Hls : oname_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In (Some n) ls -> False)
- : Named.interpf_side_conditions_gen interp_op interped_op_side_conditions ctx v
- = Some (interpf_side_conditions_gen interp_op interped_op_side_conditions e').
- Proof using ContextOk Name_dec base_type_dec.
- revert dependent ctx; revert dependent ls; induction Hwf;
- repeat first [ progress intros
- | progress subst
- | progress inversion_option
- | apply (f_equal (@Some _))
- | apply (f_equal2 (@pair _ _))
- | apply (f_equal2 and_pointed_Prop)
- | apply (f_equal (@interp_op _ _ _))
- | solve [ eauto ]
- | progress simpl in *
- | progress unfold option_map, LetIn.Let_In in *
- | progress break_innermost_match_step
- | progress break_match_hyps
- | progress destruct_head' or
- | progress inversion_prod
- | progress specialize_by_assumption
- | progress specialize_by auto using oname_list_unique_nil
- | match goal with
- | [ H : forall x, oname_list_unique ?ls -> _ |- _ ]
- => specialize (fun pf x => H x pf)
- | [ H : context[snd (split_onames _ _)] |- _ ]
- => rewrite snd_split_onames_skipn in H
- | [ H : oname_list_unique (List.skipn _ _) -> _ |- _ ]
- => specialize (fun pf => H (@oname_list_unique_skipn _ _ _ pf))
- | [ IH : forall v ls, ocompilef ?e ls = Some v -> _, H : ocompilef ?e ?ls' = Some ?v' |- _ ]
- => specialize (IH _ _ H)
- | [ IH : forall x1 x2 v ls, ocompilef (?e x1) ls = Some v -> _, H : ocompilef (?e ?x1') ?ls' = Some ?v' |- _ ]
- => specialize (fun x2 => IH _ x2 _ _ H)
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => rewrite List.in_app_iff in H
- | [ H : forall ctx, _ -> Named.interpf_side_conditions_gen _ _ _ ?e = Some _, H' : context[Named.interpf_side_conditions_gen _ _ _ ?e] |- _ ]
- => rewrite H in H' by assumption
- | [ H : forall x2 ctx, _ -> Named.interpf_side_conditions_gen _ _ _ ?e = Some _ |- Named.interpf_side_conditions_gen _ _ _ ?e = Some _ ]
- => apply H; clear H
- | [ H : forall x2, _ -> forall ctx, _ -> Named.interpf_side_conditions_gen _ _ _ ?e = Some _ |- Named.interpf_side_conditions_gen _ _ _ ?e = Some _ ]
- => apply H; clear H
- | [ H : forall x2 ctx, _ -> Named.interpf_side_conditions_gen _ _ _ ?e = Some _, H' : context[Named.interpf_side_conditions_gen _ _ _ ?e] |- _ ]
- => erewrite H in H'; clear H
- | [ H : forall x2, _ -> forall ctx, _ -> Named.interpf_side_conditions_gen _ _ _ ?e = Some _, H' : context[Named.interpf_side_conditions_gen _ _ _ ?e] |- _ ]
- => erewrite H in H'; clear H
- end ];
- repeat match goal with
- | _ => erewrite lookupb_extend by assumption
- | [ |- context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- | [ H : _ |- _ ] => erewrite H by eassumption
- | _ => progress unfold dec in *
- | _ => progress break_innermost_match_step
- | _ => progress subst
- | _ => progress destruct_head' and
- | _ => congruence
- | [ H : List.In _ (flatten_binding_list _ _) |- _ ]
- => erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) in H;
- [ | | apply path_prod_uncurried; split; [ eassumption | simpl; reflexivity ] ];
- [ | solve [ eauto using oname_list_unique_firstn, oname_list_unique_skipn ] ]
- | [ H : _ |- _ ]
- => first [ erewrite find_Name_and_val_wrong_type in H by eassumption
- | rewrite find_Name_and_val_different in H by assumption
- | rewrite snd_split_onames_skipn in H ]
- | _ => solve [ eauto using In_skipn, In_firstn
- | eapply split_onames_find_Name_Some_unique; [ | apply path_prod; simpl | ]; eauto ]
- | [ H : find_Name_and_val _ _ ?t ?n ?N ?V None = Some _, H' : List.In (Some ?n) (List.skipn _ ?ls) |- False ]
- => eapply find_Name_and_val_find_Name_Some, split_onames_find_Name_Some_unique in H;
- [ | | apply path_prod_uncurried; split; [ eassumption | simpl; reflexivity ] ];
- [ | solve [ eauto using oname_list_unique_firstn, oname_list_unique_skipn ] ]
- | [ H : List.In (existT _ ?t (?n, _)%core) ?G,
- H' : List.In (Some ?n) (List.skipn _ ?ls),
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> List.In (Some n') ?ls -> False
- |- False ]
- => apply (IH _ _ _ H); clear IH H
- | [ H : List.In (existT _ ?t (?n, _)%core) ?G,
- H' : find_Name _ ?n ?N = Some ?t',
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> List.In (Some n') ?ls -> False
- |- _ ]
- => exfalso; apply (IH _ _ _ H); clear IH H
- | [ H : List.In (existT _ ?t (?n, ?v')%core) ?G,
- H' : lookupb ?ctx ?x = Some ?v,
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> lookupb ?ctx n' = Some x'
- |- _ ]
- => assert (v = v') by (pose proof (IH _ _ _ H); congruence);
- (subst v || subst v')
- | [ H : List.In (existT _ ?t (?n, ?v')%core) ?G,
- H' : lookupb ?ctx ?x = _,
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> lookupb ?ctx n' = Some x'
- |- _ ]
- => pose proof (IH _ _ _ H); congruence
- end.
- Qed.
-
- Lemma interpf_side_conditions_ocompilef (ctx : Context) {t} (e : exprf t) e' (ls : list (option Name))
- G
- (Hwf : wff G e e')
- (HG : forall t n x, List.In (existT _ t (n, x)%core) G -> lookupb t ctx n = Some x)
- v
- (H : ocompilef e ls = Some v)
- (Hls : oname_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In (Some n) ls -> False)
- : Named.interpf_side_conditions interp_op interped_op_side_conditions ctx v
- = Some (interpf_side_conditions interp_op interped_op_side_conditions e').
- Proof using ContextOk Name_dec base_type_dec.
- unfold Named.interpf_side_conditions; erewrite interpf_side_conditions_gen_ocompilef by eassumption.
- reflexivity.
- Qed.
-
- Lemma interp_side_conditions_ocompile (ctx : Context) {t} (e : expr t) e' (ls : list (option Name))
- (Hwf : wf e e')
- f
- (Hls : oname_list_unique ls)
- (H : ocompile e ls = Some f)
- : forall v, Named.interp_side_conditions interp_op interped_op_side_conditions ctx f v
- = Some (interp_side_conditions interp_op interped_op_side_conditions e' v).
- Proof using ContextOk Name_dec base_type_dec.
- revert H; destruct Hwf;
- repeat first [ progress simpl in *
- | progress unfold option_map, Named.interp in *
- | congruence
- | progress break_innermost_match
- | progress inversion_option
- | progress subst
- | progress intros ].
- eapply interpf_side_conditions_ocompilef;
- [ eauto | | eassumption
- | inversion_prod; subst; rewrite snd_split_onames_skipn; eauto using oname_list_unique_skipn
- |intros ???; erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) by eassumption;
- let H := fresh in
- intro H; apply find_Name_and_val_find_Name_Some in H;
- eapply split_onames_find_Name_Some_unique in H; [ | eassumption.. ];
- intuition ].
- { intros ???.
- repeat first [ solve [ auto ]
- | rewrite (lookupb_extend _ _ _)
- | progress subst
- | progress break_innermost_match
- | erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) by eassumption
- | congruence
- | match goal with
- | [ |- context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- | [ H : _ |- _ ] => first [ erewrite find_Name_and_val_wrong_type in H by eassumption
- | erewrite find_Name_and_val_different in H by eassumption ]
- end
- | progress intros ]. }
- Qed.
-
- Lemma interpf_side_conditions_gen_compilef (ctx : Context) {t} (e : exprf t) e' (ls : list Name)
- G
- (Hwf : wff G e e')
- (HG : forall t n x, List.In (existT _ t (n, x)%core) G -> lookupb t ctx n = Some x)
- v
- (H : compilef e ls = Some v)
- (Hls : name_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In n ls -> False)
- : Named.interpf_side_conditions_gen interp_op interped_op_side_conditions ctx v
- = Some (interpf_side_conditions_gen interp_op interped_op_side_conditions e').
- Proof using ContextOk Name_dec base_type_dec.
- eapply interpf_side_conditions_gen_ocompilef; try eassumption.
- setoid_rewrite List.in_map_iff; intros; destruct_head' ex; destruct_head' and; inversion_option; subst.
- eauto.
- Qed.
-
- Lemma interpf_side_conditions_compilef (ctx : Context) {t} (e : exprf t) e' (ls : list Name)
- G
- (Hwf : wff G e e')
- (HG : forall t n x, List.In (existT _ t (n, x)%core) G -> lookupb t ctx n = Some x)
- v
- (H : compilef e ls = Some v)
- (Hls : name_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In n ls -> False)
- : Named.interpf_side_conditions interp_op interped_op_side_conditions ctx v
- = Some (interpf_side_conditions interp_op interped_op_side_conditions e').
- Proof using ContextOk Name_dec base_type_dec.
- unfold Named.interpf_side_conditions; erewrite interpf_side_conditions_gen_compilef by eassumption.
- reflexivity.
- Qed.
-
- Lemma interp_side_conditions_compile (ctx : Context) {t} (e : expr t) e' (ls : list Name)
- (Hwf : wf e e')
- f
- (Hls : name_list_unique ls)
- (H : compile e ls = Some f)
- : forall v, Named.interp_side_conditions interp_op interped_op_side_conditions ctx f v
- = Some (interp_side_conditions interp_op interped_op_side_conditions e' v).
- Proof using ContextOk Name_dec base_type_dec. eapply interp_side_conditions_ocompile; eassumption. Qed.
-
- Lemma InterpSideConditions_compile {t} (e : Expr t) (ls : list Name)
- (Hwf : Wf e)
- f
- (Hls : name_list_unique ls)
- (H : compile (e _) ls = Some f)
- : forall v, Named.InterpSideConditions (Context:=Context) interp_op interped_op_side_conditions f v
- = Some (InterpSideConditions interp_op interped_op_side_conditions e v).
- Proof using ContextOk Name_dec base_type_dec. eapply interp_side_conditions_compile; eauto. Qed.
-End language.
diff --git a/src/Compilers/Named/CompileProperties.v b/src/Compilers/Named/CompileProperties.v
deleted file mode 100644
index 9803946b2..000000000
--- a/src/Compilers/Named/CompileProperties.v
+++ /dev/null
@@ -1,74 +0,0 @@
-Require Import Coq.omega.Omega.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.Named.NameUtilProperties.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.CountLets.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.ListUtil.
-
-Local Notation eta x := (fst x, snd x).
-
-Local Open Scope ctype_scope.
-Local Open Scope nexpr_scope.
-Local Open Scope expr_scope.
-Section language.
- Context (base_type_code : Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (Name : Type)
- (dummy : base_type_code -> Name).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprft := (@exprf base_type_code op (fun _ => unit)).
- Local Notation exprt := (@expr base_type_code op (fun _ => unit)).
- Local Notation exprf := (@exprf base_type_code op (fun _ => Name)).
- Local Notation expr := (@expr base_type_code op (fun _ => Name)).
- Local Notation nexprf := (@Named.exprf base_type_code op Name).
- Local Notation nexpr := (@Named.expr base_type_code op Name).
-
- Lemma compilef_count_let_bindersf_enough {t}
- (e1 : exprf t) (e2 : exprft t)
- G
- (Hwf : wff G e1 e2)
- : forall (ls1 : list Name)
- (He : compilef e1 ls1 <> None)
- (ls2 : list Name)
- (Hls : List.length ls2 >= count_let_bindersf dummy e1),
- compilef e1 ls2 <> None.
- Proof.
- unfold compilef; induction Hwf;
- repeat first [ progress simpl in *
- | progress cbv [option_map] in *
- | progress subst
- | progress inversion_prod
- | congruence
- | omega
- | progress break_innermost_match_step
- | progress break_match_hyps
- | progress intros
- | progress specialize_by congruence
- | match goal with
- | [ H : forall ls1, ocompilef ?e _ <> None -> _, H' : ocompilef ?e (List.map _ ?ls') = Some _ |- _ ]
- => specialize (H ls'); rewrite H' in H
- | [ H : forall ls1, ocompilef ?e _ <> None -> _, H' : ocompilef ?e nil = Some _ |- _ ]
- => specialize (H nil); simpl in H; rewrite H' in H
- | [ H : forall v ls1, ocompilef (?e v) _ <> None -> _, H' : ocompilef (?e ?v') _ = Some _ |- _ ]
- => specialize (H v')
- | [ H : forall ls1, List.length ls1 >= ?k -> _, H' : List.length _ >= ?k |- _ ]
- => specialize (H _ H')
- | [ H : context[snd (split_onames _ _)] |- _ ]
- => rewrite snd_split_onames_skipn in H
- | [ H : context[List.skipn _ (List.map _ _)] |- _]
- => rewrite skipn_map in H
- | [ H : fst (split_onames ?t (List.map _ ?ls)) = None |- _ ]
- => rewrite split_onames_split_names in H
- | [ H : fst (split_names _ _) = None |- _ ]
- => apply length_fst_split_names_None_iff in H
- end ].
- Abort.
-End language.
diff --git a/src/Compilers/Named/CompileWf.v b/src/Compilers/Named/CompileWf.v
deleted file mode 100644
index 3ccff8d58..000000000
--- a/src/Compilers/Named/CompileWf.v
+++ /dev/null
@@ -1,254 +0,0 @@
-(** * PHOAS → Named Representation of Gallina *)
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.NameUtil.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.Named.NameUtilProperties.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Local Open Scope ctype_scope.
-Local Open Scope nexpr_scope.
-Local Open Scope expr_scope.
-Section language.
- Context {base_type_code var} {var' : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {base_type_dec : DecidableRel (@eq base_type_code)}
- {Name_dec : DecidableRel (@eq Name)}
- {Context : @Context base_type_code Name var}
- {ContextOk : ContextOk Context}
- {make_var' : forall t, var t -> var' t} (* probably only needed because I was clumsy in the proof method *).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op (fun _ => Name)).
- Local Notation expr := (@expr base_type_code op (fun _ => Name)).
- Local Notation wff := (@wff base_type_code op (fun _ => Name) var').
- Local Notation wf := (@wf base_type_code op (fun _ => Name) var').
- Local Notation nexprf := (@Named.exprf base_type_code op Name).
- Local Notation nexpr := (@Named.expr base_type_code op Name).
- Local Notation nwff := (@Named.wff base_type_code Name op var Context).
- Local Notation nwf := (@Named.wf base_type_code Name op var Context).
- Local Notation ocompilef := (@ocompilef base_type_code op Name).
- Local Notation ocompile := (@ocompile base_type_code op Name).
- Local Notation compilef := (@compilef base_type_code op Name).
- Local Notation compile := (@compile base_type_code op Name).
-
- Local Ltac finish_with_var' :=
- (* FIXME: clean this up *)
- lazymatch goal with
- | [ H : find_Name_and_val _ _ ?b ?n ?i ?v None = Some _
- |- find_Name_and_val _ _ ?b ?n ?i _ None <> None ]
- => (is_evar v;
- let T := type of v in
- let H := fresh in
- assert (H : { k : T | k = v }));
- [
- | let H' := fresh in
- let H'' := fresh in
- intro H';
- let T := match type of H with ?T = _ => T end in
- assert (H'' : T <> None) by congruence; apply H''; revert H';
- rewrite <- !find_Name_and_val_None_iff;
- tauto ]
- end;
- match goal with
- | [ v : interp_flat_type _ ?t |- @sig (interp_flat_type _ ?t) _ ]
- => exists (SmartMap.SmartVarfMap make_var' v)
- end;
- reflexivity.
-
- Lemma wff_ocompilef (ctx : Context) G
- (HG : forall t n v,
- List.In (existT _ t (n, v)%core) G -> lookupb t ctx n <> None)
- {t} (e : exprf t) e' (ls : list (option Name))
- (Hwf : wff G e e')
- v
- (H : ocompilef e ls = Some v)
- (Hls : oname_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In (Some n) ls -> False)
- : prop_of_option (nwff ctx v).
- Proof using ContextOk Name_dec base_type_dec make_var'.
- revert dependent ctx; revert dependent ls; induction Hwf;
- repeat first [ progress intros
- | progress subst
- | progress inversion_option
- | solve [ auto ]
- | progress simpl in *
- | progress unfold option_map in *
- | progress break_innermost_match_step
- | progress break_match_hyps
- | progress autorewrite with push_prop_of_option in *
- | progress specialize_by tauto
- | progress specialize_by auto using oname_list_unique_nil
- | solve [ unfold not in *; eauto using In_skipn, oname_list_unique_firstn, oname_list_unique_skipn ]
- | progress destruct_head' or
- | match goal with
- | [ IH : forall v ls, ocompilef ?e ls = Some v -> _, H : ocompilef ?e ?ls' = Some ?v' |- _ ]
- => specialize (IH _ _ H)
- | [ IH : forall x1 x2 v ls, ocompilef (?e2 x1) ls = Some v -> _, H : ocompilef (?e2 ?x1') ?ls' = Some ?v' |- _ ]
- => specialize (fun x2 => IH _ x2 _ _ H)
- | [ HG : forall t n v, List.In _ _ -> _ = Some _, H : _ = None |- _ ]
- => erewrite HG in H by eassumption
- | [ |- _ /\ _ ] => split
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => setoid_rewrite List.in_app_iff in H
- | [ H : split_onames _ _ = (_, _)%core |- _ ]
- => pose proof (f_equal (@fst _ _) H);
- pose proof (f_equal (@snd _ _) H);
- clear H
- | [ H : context[snd (split_onames _ _)] |- _ ]
- => rewrite snd_split_onames_skipn in H
- | [ H : forall a, (forall x y z, _ \/ _ -> _) -> _ |- _ ]
- => specialize (fun a pf1 pf2
- => H a (fun x y z pf
- => match pf with
- | or_introl pf' => pf1 x y z pf'
- | or_intror pf' => pf2 x y z pf'
- end))
- | [ H : forall a b, (forall x y z, _ \/ _ -> _) -> _ |- _ ]
- => specialize (fun a b pf1 pf2
- => H a b (fun x y z pf
- => match pf with
- | or_introl pf' => pf1 x y z pf'
- | or_intror pf' => pf2 x y z pf'
- end))
- | [ H : forall a b c, (forall x y z, _ \/ _ -> _) -> _ |- _ ]
- => specialize (fun a b c pf1 pf2
- => H a b c (fun x y z pf
- => match pf with
- | or_introl pf' => pf1 x y z pf'
- | or_intror pf' => pf2 x y z pf'
- end))
- | [ H : forall a b c d, (forall x y z, _ \/ _ -> _) -> _ |- _ ]
- => specialize (fun a b c d pf1 pf2
- => H a b c d (fun x y z pf
- => match pf with
- | or_introl pf' => pf1 x y z pf'
- | or_intror pf' => pf2 x y z pf'
- end))
- | [ H : _ |- _ ]
- => progress rewrite ?firstn_nil, ?skipn_nil, ?skipn_skipn in H
- | [ H : List.In ?x (List.firstn ?a (List.skipn ?b ?ls)), H' : List.In ?x (List.skipn (?b + ?a) ?ls) |- False ]
- => rewrite firstn_skipn_add in H; apply In_skipn in H
- | [ H : ?T |- prop_of_option (nwff _ ?v) ]
- => eapply H; clear H
- end ];
- repeat match goal with
- | _ => erewrite lookupb_extend by assumption
- | [ |- context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- | [ H : _ |- _ ] => erewrite H by eassumption
- | _ => progress unfold dec in *
- | _ => progress break_innermost_match_step
- | _ => progress subst
- | _ => progress destruct_head' and
- | _ => congruence
- | [ H : List.In _ (flatten_binding_list _ _) |- _ ]
- => erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) in H;
- [ | | apply path_prod_uncurried; split; [ eassumption | simpl; reflexivity ] ];
- [ | solve [ eauto using oname_list_unique_firstn, oname_list_unique_skipn ] ]
- | [ H : _ |- _ ]
- => first [ erewrite find_Name_and_val_wrong_type in H by eassumption
- | rewrite find_Name_and_val_different in H by assumption
- | rewrite snd_split_onames_skipn in H ]
- | _ => solve [ eauto using In_skipn, In_firstn
- | eapply split_onames_find_Name_Some_unique; [ | apply path_prod; simpl | ]; eauto ]
- | [ H : find_Name_and_val _ _ ?t ?n ?N ?V None = Some _, H' : List.In (Some ?n) (List.skipn _ ?ls) |- False ]
- => eapply find_Name_and_val_find_Name_Some, split_onames_find_Name_Some_unique in H;
- [ | | apply path_prod_uncurried; split; [ eassumption | simpl; reflexivity ] ];
- [ | solve [ eauto using oname_list_unique_firstn, oname_list_unique_skipn ] ]
- | [ H : List.In (existT _ ?t (?n, _)%core) ?G,
- H' : List.In (Some ?n) (List.skipn _ ?ls),
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> List.In (Some n') ?ls -> False
- |- False ]
- => apply (IH _ _ _ H); clear IH H
- | [ H : List.In (existT _ ?t (?n, _)%core) ?G,
- H' : find_Name _ ?n ?N = Some ?t',
- IH : forall t' n' x', List.In (existT _ t' (n', x')%core) ?G -> List.In (Some n') ?ls -> False
- |- _ ]
- => exfalso; apply (IH _ _ _ H); clear IH H
- end.
- finish_with_var'.
- Qed.
-
- Lemma wf_ocompile (ctx : Context) {t} (e : expr t) e' (ls : list (option Name))
- (Hwf : wf e e')
- f
- (Hls : oname_list_unique ls)
- (H : ocompile e ls = Some f)
- : nwf ctx f.
- Proof using ContextOk Name_dec base_type_dec make_var'.
- revert H; destruct Hwf;
- repeat first [ progress simpl in *
- | progress unfold option_map, Named.interp in *
- | congruence
- | progress break_innermost_match
- | progress inversion_option
- | progress subst
- | progress intros ].
- intro; simpl.
- eapply wff_ocompilef;
- [ | solve [ eauto ] | eassumption
- | inversion_prod; subst; rewrite snd_split_onames_skipn; eauto using oname_list_unique_skipn
- | intros ???; erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) by eassumption;
- let H := fresh in
- intro H; apply find_Name_and_val_find_Name_Some in H;
- eapply split_onames_find_Name_Some_unique in H; [ | eassumption.. ];
- intuition ].
- { intros ???.
- repeat first [ solve [ auto ]
- | rewrite (lookupb_extend _ _ _)
- | progress subst
- | progress break_innermost_match
- | erewrite <- (flatten_binding_list_find_Name_and_val_unique _ _) by eassumption
- | congruence
- | eassumption
- | match goal with
- | [ |- context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- | [ H : _ |- _ ] => first [ erewrite find_Name_and_val_wrong_type in H by eassumption
- | erewrite find_Name_and_val_different in H by eassumption ]
- end
- | progress intros ].
- finish_with_var'. }
- Qed.
-
- Lemma wff_compilef (ctx : Context) {t} (e : exprf t) e' (ls : list Name)
- G
- (Hwf : wff G e e')
- (HG : forall t n x, List.In (existT _ t (n, x)%core) G -> lookupb t ctx n <> None)
- v
- (H : compilef e ls = Some v)
- (Hls : name_list_unique ls)
- (HGls : forall t n x, List.In (existT _ t (n, x)%core) G -> List.In n ls -> False)
- : prop_of_option (nwff ctx v).
- Proof using ContextOk Name_dec base_type_dec make_var'.
- eapply wff_ocompilef; try eassumption.
- setoid_rewrite List.in_map_iff; intros; destruct_head' ex; destruct_head' and; inversion_option; subst.
- eauto.
- Qed.
-
- Lemma wf_compile (ctx : Context) {t} (e : expr t) e' (ls : list Name)
- (Hwf : wf e e')
- f
- (Hls : name_list_unique ls)
- (H : compile e ls = Some f)
- : nwf ctx f.
- Proof using ContextOk Name_dec base_type_dec make_var'. eapply wf_ocompile; eassumption. Qed.
-End language.
diff --git a/src/Compilers/Named/Context.v b/src/Compilers/Named/Context.v
deleted file mode 100644
index fd84e5b94..000000000
--- a/src/Compilers/Named/Context.v
+++ /dev/null
@@ -1,65 +0,0 @@
-(** * Contexts for Named Representation of Gallina *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Util.Notations.
-
-Record Context {base_type_code} (Name : Type) (var : base_type_code -> Type) :=
- { ContextT : Type;
- lookupb : forall {t : base_type_code}, ContextT -> Name -> option (var t);
- extendb : forall {t : base_type_code}, ContextT -> Name -> var t -> ContextT;
- removeb : base_type_code -> ContextT -> Name -> ContextT;
- empty : ContextT }.
-Coercion ContextT : Context >-> Sortclass.
-Arguments ContextT {_ _ _ _}, {_ _ _} _.
-Arguments lookupb {_ _ _ _ _} _ _, {_ _ _ _} _ _ _.
-Arguments extendb {_ _ _ _} [_] _ _ _.
-Arguments removeb {_ _ _ _} _ _ _.
-Arguments empty {_ _ _ _}.
-
-Section language.
- Context {base_type_code : Type}
- {Name : Type}
- {var : base_type_code -> Type}
- {Context : Context Name var}.
-
- Local Notation flat_type := (flat_type base_type_code).
-
- Fixpoint extend {t : flat_type} (ctx : Context)
- (n : interp_flat_type (fun _ => Name) t) (v : interp_flat_type var t)
- : Context
- := match t return interp_flat_type (fun _ => Name) t -> interp_flat_type var t -> Context with
- | Tbase t => fun n v => extendb ctx n v
- | Unit => fun _ _ => ctx
- | Prod A B => fun n v : interp_flat_type _ A * interp_flat_type _ B
- => let ctx := @extend A ctx (fst n) (fst v) in
- let ctx := @extend B ctx (snd n) (snd v) in
- ctx
- end n v.
-
- Fixpoint remove {t : flat_type} (ctx : Context)
- (n : interp_flat_type (fun _ => Name) t)
- : Context
- := match t return interp_flat_type (fun _ => Name) t -> Context with
- | Tbase t => fun n => removeb t ctx n
- | Unit => fun _ => ctx
- | Prod A B => fun n : interp_flat_type _ A * interp_flat_type _ B
- => let ctx := @remove A ctx (fst n) in
- let ctx := @remove B ctx (snd n) in
- ctx
- end n.
-
- Definition lookup {t} (ctx : Context)
- : interp_flat_type (fun _ => Name) t -> option (interp_flat_type var t)
- := smart_interp_flat_map
- (g := fun t => option (interp_flat_type var t))
- (fun t v => lookupb ctx v)
- (Some tt)
- (fun A B x y => match x, y with
- | Some x', Some y' => Some (x', y')%core
- | _, _ => None
- end).
-End language.
-
-Global Arguments extend {_ _ _ _} {_} ctx _ _.
-Global Arguments remove {_ _ _ _} {_} ctx _.
-Global Arguments lookup {_ _ _ _} {_} ctx _, {_ _ _ _} _ ctx _.
diff --git a/src/Compilers/Named/ContextDefinitions.v b/src/Compilers/Named/ContextDefinitions.v
deleted file mode 100644
index 1b57e5b51..000000000
--- a/src/Compilers/Named/ContextDefinitions.v
+++ /dev/null
@@ -1,67 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Decidable.
-
-Section with_context.
- Context {base_type_code Name var} (Context : @Context base_type_code Name var)
- (base_type_code_dec : DecidableRel (@eq base_type_code))
- (Name_dec : DecidableRel (@eq Name)).
-
- Fixpoint find_Name n
- {T : flat_type base_type_code}
- : interp_flat_type (fun _ => Name) T -> option base_type_code
- := match T with
- | Tbase t' => fun n' : Name => if dec (n = n') then Some t' else None
- | Unit => fun _ => None
- | Prod A B
- => fun ab : interp_flat_type _ A * interp_flat_type _ B
- => match @find_Name n B (snd ab), @find_Name n A (fst ab) with
- | Some tb, _ => Some tb
- | None, Some ta => Some ta
- | None, None => None
- end
- end.
-
- Fixpoint find_Name_and_val {var'} t (n : Name)
- {T : flat_type base_type_code}
- : interp_flat_type (fun _ => Name) T -> interp_flat_type var' T -> option (var' t) -> option (var' t)
- := match T with
- | Tbase t' => fun (n' : Name) v default
- => if dec (n = n')
- then cast_if_eq t' t v
- else default
- | Unit => fun _ _ default => default
- | Prod A B
- => fun (ab : interp_flat_type _ A * interp_flat_type _ B)
- (a'b' : interp_flat_type _ A * interp_flat_type _ B)
- default
- => @find_Name_and_val
- var' t n B (snd ab) (snd a'b')
- (@find_Name_and_val
- var' t n A (fst ab) (fst a'b')
- default)
- end.
-
- Class ContextOk :=
- { lookupb_extendb_same
- : forall (ctx : Context) n t v, lookupb t (extendb ctx n (t:=t) v) n = Some v;
- lookupb_extendb_different
- : forall (ctx : Context) n n' t t' v, n <> n' -> lookupb t' (extendb ctx n (t:=t) v) n'
- = lookupb t' ctx n';
- lookupb_extendb_wrong_type
- : forall (ctx : Context) n t t' v, t <> t' -> lookupb t' (extendb ctx n (t:=t) v) n = None;
- lookupb_removeb_different
- : forall (ctx : Context) n n' t t', n <> n' -> lookupb t' (removeb t ctx n) n'
- = lookupb t' ctx n';
- lookupb_removeb_same
- : forall (ctx : Context) n t t', lookupb t' (removeb t ctx n) n = None;
- lookupb_empty
- : forall n t, lookupb t (@empty _ _ _ Context) n = None;
-
- lookupb_unique_type
- : forall (ctx : Context) n t t', lookupb t' ctx n <> None -> lookupb t ctx n <> None -> t = t' }.
-
- Definition context_equiv (a b : Context)
- := forall n t, lookupb t a n = lookupb t b n.
-End with_context.
diff --git a/src/Compilers/Named/ContextOn.v b/src/Compilers/Named/ContextOn.v
deleted file mode 100644
index 17d8690fc..000000000
--- a/src/Compilers/Named/ContextOn.v
+++ /dev/null
@@ -1,24 +0,0 @@
-(** * Transfer [Context] across an injection *)
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-
-Section language.
- Context {base_type_code Name1 Name2 : Type}
- (f : Name2 -> Name1)
- (f_inj : forall x y, f x = f y -> x = y)
- {var : base_type_code -> Type}.
-
- Definition ContextOn (Ctx : Context Name1 var) : Context Name2 var
- := {| ContextT := Ctx;
- lookupb t ctx n := lookupb t ctx (f n);
- extendb t ctx n v := extendb ctx (f n) v;
- removeb t ctx n := removeb t ctx (f n);
- empty := empty |}.
-
- Lemma ContextOnOk {Ctx} (COk : ContextOk Ctx) : ContextOk (ContextOn Ctx).
- Proof.
- unfold ContextOn in *; split; intros; try eapply COk; eauto.
- Qed.
-End language.
-
-Arguments ContextOnOk {_ _ _ f} f_inj {_ _} COk.
diff --git a/src/Compilers/Named/ContextProperties.v b/src/Compilers/Named/ContextProperties.v
deleted file mode 100644
index 782fc9a54..000000000
--- a/src/Compilers/Named/ContextProperties.v
+++ /dev/null
@@ -1,181 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.Tactics.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Section with_context.
- Context {base_type_code}
- (base_type_code_dec : DecidableRel (@eq base_type_code))
- {Name}
- (Name_dec : DecidableRel (@eq Name))
- {var} (Context : @Context base_type_code Name var)
- (ContextOk : ContextOk Context).
-
- Local Notation find_Name := (@find_Name base_type_code Name Name_dec).
- Local Notation find_Name_and_val := (@find_Name_and_val base_type_code Name base_type_code_dec Name_dec).
-
- Lemma lookupb_eq_cast
- : forall (ctx : Context) n t t' (pf : t = t'),
- lookupb t' ctx n = option_map (fun v => eq_rect _ var v _ pf) (lookupb t ctx n).
- Proof.
- intros; subst; edestruct lookupb; reflexivity.
- Defined.
- Lemma lookupb_extendb_eq
- : forall (ctx : Context) n t t' (pf : t = t') v,
- lookupb t' (extendb ctx n (t:=t) v) n = Some (eq_rect _ var v _ pf).
- Proof.
- intros; subst; apply lookupb_extendb_same; assumption.
- Defined.
-
- Lemma lookupb_extendb_full (ctx : Context) n n' t t' v
- : lookupb t' (extendb ctx n (t:=t) v) n'
- = match dec (n = n'), dec (t = t') with
- | left _, left pf
- => Some (eq_rect _ var v _ pf)
- | left _, _
- => None
- | right _, _
- => lookupb t' ctx n'
- end.
- Proof using ContextOk.
- break_innermost_match; subst; simpl.
- { apply lookupb_extendb_same; assumption. }
- { apply lookupb_extendb_wrong_type; assumption. }
- { apply lookupb_extendb_different; assumption. }
- Qed.
-
- Lemma lookupb_extend (ctx : Context)
- T N t n v
- : lookupb t (extend ctx N (t:=T) v) n
- = find_Name_and_val t n N v (lookupb t ctx n).
- Proof using ContextOk. revert ctx; induction T; t. Qed.
-
- Lemma lookupb_remove (ctx : Context)
- T N t n
- : lookupb t (remove ctx N) n
- = match @find_Name n T N with
- | Some _ => None
- | None => lookupb t ctx n
- end.
- Proof using ContextOk. revert ctx; induction T; t. Qed.
-
- Lemma lookupb_remove_not_in (ctx : Context)
- T N t n
- (H : @find_Name n T N = None)
- : lookupb t (remove ctx N) n = lookupb t ctx n.
- Proof using ContextOk. rewrite lookupb_remove, H; reflexivity. Qed.
-
- Lemma lookupb_remove_in (ctx : Context)
- T N t n
- (H : @find_Name n T N <> None)
- : lookupb t (remove ctx N) n = None.
- Proof using ContextOk. rewrite lookupb_remove; break_match; congruence. Qed.
-
- Lemma find_Name_and_val_Some_None
- {var' var''}
- {t n T N}
- {x : interp_flat_type var' T}
- {y : interp_flat_type var'' T}
- {default default' v}
- (H0 : @find_Name_and_val var' t n T N x default = Some v)
- (H1 : @find_Name_and_val var'' t n T N y default' = None)
- : default = Some v /\ default' = None.
- Proof using Type.
- revert dependent default; revert dependent default'; induction T; t.
- Qed.
-
- Lemma find_Name_and_val_default_to_None
- {var'}
- {t n T N}
- {x : interp_flat_type var' T}
- {default}
- (H : @find_Name n T N <> None)
- : @find_Name_and_val var' t n T N x default
- = @find_Name_and_val var' t n T N x None.
- Proof using Type. revert default; induction T; t. Qed.
- Hint Rewrite @find_Name_and_val_default_to_None using congruence : ctx_db.
-
- Lemma find_Name_and_val_different
- {var'}
- {t n T N}
- {x : interp_flat_type var' T}
- {default}
- (H : @find_Name n T N = None)
- : @find_Name_and_val var' t n T N x default = default.
- Proof using Type. revert default; induction T; t. Qed.
- Hint Rewrite @find_Name_and_val_different using assumption : ctx_db.
-
- Lemma find_Name_and_val_wrong_type_iff
- {var'}
- {t t' n T N}
- {x : interp_flat_type var' T}
- {default}
- (H : @find_Name n T N = Some t')
- : t <> t'
- <-> @find_Name_and_val var' t n T N x default = None.
- Proof using Type. split; revert default; induction T; t. Qed.
- Lemma find_Name_and_val_wrong_type
- {var'}
- {t t' n T N}
- {x : interp_flat_type var' T}
- {default}
- (H : @find_Name n T N = Some t')
- (Ht : t <> t')
- : @find_Name_and_val var' t n T N x default = None.
- Proof using Type. eapply find_Name_and_val_wrong_type_iff; eassumption. Qed.
- Hint Rewrite @find_Name_and_val_wrong_type using congruence : ctx_db.
-
- Lemma find_Name_find_Name_and_val_wrong {var' n t' T V N}
- : find_Name n N = Some t'
- -> @find_Name_and_val var' t' n T N V None = None
- -> False.
- Proof using Type. induction T; t. Qed.
-
- Lemma find_Name_and_val_None_iff
- {var'}
- {t n T N}
- {x : interp_flat_type var' T}
- {default}
- : (@find_Name n T N <> Some t
- /\ (@find_Name n T N <> None \/ default = None))
- <-> @find_Name_and_val var' t n T N x default = None.
- Proof using Type.
- destruct (@find_Name n T N) eqn:?; unfold not; t;
- try solve [ eapply find_Name_and_val_wrong_type; [ eassumption | congruence ]
- | eapply find_Name_find_Name_and_val_wrong; eassumption
- | left; congruence ].
- Qed.
-
- Lemma find_Name_and_val_split
- {var' t n T N V default}
- : @find_Name_and_val var' t n T N V default
- = match @find_Name n T N with
- | Some t' => if dec (t = t')
- then @find_Name_and_val var' t n T N V None
- else None
- | None => default
- end.
- Proof using Type.
- t; erewrite find_Name_and_val_wrong_type by solve [ eassumption | congruence ]; reflexivity.
- Qed.
- Lemma find_Name_and_val_find_Name_Some
- {var' t n T N V v}
- (H : @find_Name_and_val var' t n T N V None = Some v)
- : @find_Name n T N = Some t.
- Proof using Type.
- rewrite find_Name_and_val_split in H; break_match_hyps; subst; congruence.
- Qed.
-End with_context.
-
-Ltac find_Name_and_val_default_to_None_step :=
- match goal with
- | [ H : context[find_Name_and_val ?tdec ?ndec _ _ _ _ ?default] |- _ ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default)) in H
- | [ |- context[find_Name_and_val ?tdec ?ndec _ _ _ _ ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- end.
-Ltac find_Name_and_val_default_to_None := repeat find_Name_and_val_default_to_None_step.
diff --git a/src/Compilers/Named/ContextProperties/NameUtil.v b/src/Compilers/Named/ContextProperties/NameUtil.v
deleted file mode 100644
index 320d910f0..000000000
--- a/src/Compilers/Named/ContextProperties/NameUtil.v
+++ /dev/null
@@ -1,161 +0,0 @@
-Require Import Coq.omega.Omega.
-Require Import Crypto.Util.FixCoqMistakes.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.Named.NameUtilProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.Tactics.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-
-Section with_context.
- Context {base_type_code Name var} (Context : @Context base_type_code Name var)
- (base_type_code_dec : DecidableRel (@eq base_type_code))
- (Name_dec : DecidableRel (@eq Name))
- (ContextOk : ContextOk Context).
-
- Local Notation find_Name := (@find_Name base_type_code Name Name_dec).
- Local Notation find_Name_and_val := (@find_Name_and_val base_type_code Name base_type_code_dec Name_dec).
-
- Hint Rewrite (@find_Name_and_val_default_to_None _ base_type_code_dec _ Name_dec) using congruence : ctx_db.
- Hint Rewrite (@find_Name_and_val_different _ base_type_code_dec _ Name_dec) using assumption : ctx_db.
- Hint Rewrite (@find_Name_and_val_wrong_type _ base_type_code_dec _ Name_dec) using congruence : ctx_db.
- Hint Rewrite (@snd_split_onames_skipn base_type_code Name) : ctx_db.
-
- Local Ltac misc_oname_t_step :=
- match goal with
- | [ H : oname_list_unique (List.skipn _ _) -> _ |- _ ]
- => specialize (fun pf => H (@oname_list_unique_skipn _ _ _ pf))
- | [ H : ((_, _) = (_, _))%core -> _ |- _ ]
- => specialize (fun a b => H (f_equal2 (@pair _ _) a b))
- | [ H : ?x = (_,_)%core -> _ |- _ ]
- => rewrite (surjective_pairing x) in H;
- specialize (fun a b => H (f_equal2 (@pair _ _) a b))
- end.
-
- Lemma split_onames_find_Name
- {n T N ls ls'}
- (H : split_onames _ ls = (Some N, ls')%core)
- : (exists t, @find_Name n T N = Some t)
- <-> List.In (Some n) (List.firstn (CountLets.count_pairs T) ls).
- Proof using Type.
- revert dependent ls; intro ls; revert ls ls'; induction T; intros ls ls' H;
- [ | | specialize (IHT1 (fst N) ls (snd (split_onames T1 ls)));
- specialize (IHT2 (snd N) (snd (split_onames T1 ls)) (snd (split_onames (T1 * T2) ls))) ];
- repeat first [ misc_oname_t_step
- | t_step
- | progress split_iff
- | progress specialize_by (eexists; eauto)
- | solve [ eauto using In_skipn, In_firstn ]
- | match goal with
- | [ H : List.In ?x (List.firstn ?n ?ls) |- List.In ?x (List.firstn (?n + ?m) ?ls) ]
- => apply (In_firstn n); rewrite firstn_firstn by omega
- | [ H : _ |- _ ] => first [ rewrite firstn_skipn_add in H
- | rewrite firstn_firstn in H by omega ]
- | [ H : List.In ?x' (List.firstn (?n + ?m) ?ls) |- List.In ?x' (List.firstn ?m (List.skipn ?n ?ls)) ]
- => apply (In_firstn_skipn_split n) in H
- end ].
- Qed.
-
- Lemma split_onames_find_Name_Some_unique_iff
- {n T N ls ls'}
- (Hls : oname_list_unique ls)
- (H : split_onames _ ls = (Some N, ls')%core)
- : (exists t, @find_Name n T N = Some t)
- <-> List.In (Some n) ls /\ ~List.In (Some n) ls'.
- Proof using Type.
- rewrite (split_onames_find_Name (ls':=ls') (ls:=ls)) by assumption.
- rewrite (surjective_pairing (split_onames _ _)) in H.
- rewrite fst_split_onames_firstn, snd_split_onames_skipn in H.
- inversion_prod; subst.
- split; [ split | intros [? ?] ]; eauto using In_firstn, oname_list_unique_specialize.
- match goal with
- | [ H : List.In (Some _) ?ls |- _ ]
- => is_var ls;
- eapply In_firstn_skipn_split in H; destruct_head' or; eauto; exfalso; eauto
- end.
- Qed.
-
- Lemma split_onames_find_Name_Some_unique
- {t n T N ls ls'}
- (Hls : oname_list_unique ls)
- (H : split_onames _ ls = (Some N, ls')%core)
- (Hfind : @find_Name n T N = Some t)
- : List.In (Some n) ls /\ ~List.In (Some n) ls'.
- Proof using Type.
- eapply split_onames_find_Name_Some_unique_iff; eauto.
- Qed.
-
- Lemma flatten_binding_list_find_Name_and_val_unique
- {var' t n T N V v ls ls'}
- (Hls : oname_list_unique ls)
- (H : split_onames _ ls = (Some N, ls')%core)
- : @find_Name_and_val var' t n T N V None = Some v
- <-> List.In (existT (fun t => (Name * var' t)%type) t (n, v)) (Wf.flatten_binding_list N V).
- Proof using Type.
- revert dependent ls; intro ls; revert ls ls'; induction T; intros ls ls' Hls H;
- [ | | specialize (IHT1 (fst N) (fst V) ls (snd (split_onames T1 ls)));
- specialize (IHT2 (snd N) (snd V) (snd (split_onames T1 ls)) (snd (split_onames (T1 * T2) ls))) ];
- repeat first [ find_Name_and_val_default_to_None_step
- | progress simpl in *
- | rewrite List.in_app_iff
- | misc_oname_t_step
- | t_step
- | progress split_iff
- | lazymatch goal with
- | [ H : find_Name ?n ?x = Some ?t, H' : find_Name_and_val ?t' ?n ?X ?V None = Some ?v |- _ ]
- => apply find_Name_and_val_find_Name_Some in H'
- | [ H : find_Name ?n ?x = Some ?t, H' : find_Name ?n ?x' = Some ?t' |- _ ]
- => let apply_in_tac H :=
- (eapply split_onames_find_Name_Some_unique in H;
- [ | | apply path_prod_uncurried; split; [ eassumption | simpl; reflexivity ] ];
- [ | solve [ eauto using oname_list_unique_firstn, oname_list_unique_skipn ] ]) in
- first [ constr_eq x x'; fail 1
- | apply_in_tac H; apply_in_tac H' ]
- end ].
- Qed.
-
- Lemma fst_split_mnames__flatten_binding_list__find_Name
- (MName : Type) (force : MName -> option Name)
- {var' t n T N V v} {ls : list MName}
- (Hs : fst (split_mnames force T ls) = Some N)
- (HN : List.In (existT _ t (n, v)%core) (Wf.flatten_binding_list (var2:=var') N V))
- : find_Name n N = Some t.
- Proof.
- revert dependent ls; induction T;
- [ | | specialize (IHT1 (fst N) (fst V));
- specialize (IHT2 (snd N) (snd V)) ];
- repeat first [ misc_oname_t_step
- | t_step
- | match goal with
- | [ H : _ |- _ ] => first [ rewrite snd_split_mnames_skipn in H
- | rewrite List.in_app_iff in H ]
- | [ H : context[fst (split_mnames _ _ ?ls)] |- _ ]
- => is_var ls; rewrite (@fst_split_mnames_firstn _ _ _ _ _ ls) in H
- end ].
- Abort.
-
- Lemma fst_split_mnames__find_Name__flatten_binding_list
- (MName : Type) (force : MName -> option Name)
- {var' t n T N V v default} {ls : list MName}
- (Hs : fst (split_mnames force T ls) = Some N)
- (Hfind : find_Name n N = Some t)
- (HN : List.In (existT _ t (n, v)%core) (Wf.flatten_binding_list N V))
- : @find_Name_and_val var' t n T N V default = Some v.
- Proof.
- revert default; revert dependent ls; induction T;
- [ | | specialize (IHT1 (fst N) (fst V));
- specialize (IHT2 (snd N) (snd V)) ];
- repeat first [ find_Name_and_val_default_to_None_step
- | rewrite List.in_app_iff in *
- | t_step ].
- Abort.
-End with_context.
diff --git a/src/Compilers/Named/ContextProperties/Proper.v b/src/Compilers/Named/ContextProperties/Proper.v
deleted file mode 100644
index 0c583e446..000000000
--- a/src/Compilers/Named/ContextProperties/Proper.v
+++ /dev/null
@@ -1,142 +0,0 @@
-Require Import Coq.Classes.Morphisms.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Section with_context.
- Context {base_type_code}
- {base_type_code_dec : DecidableRel (@eq base_type_code)}
- {Name}
- {Name_dec : DecidableRel (@eq Name)}
- {var} {Context : @Context base_type_code Name var}
- {ContextOk : ContextOk Context}.
-
- Local Notation context_equiv := (@context_equiv base_type_code Name var Context).
- Local Notation extendb := (@extendb base_type_code Name var Context).
- Local Notation removeb := (@removeb base_type_code Name var Context).
- Local Notation lookupb := (@lookupb base_type_code Name var Context).
- Local Notation extend := (@extend base_type_code Name var Context).
- Local Notation remove := (@remove base_type_code Name var Context).
- Local Notation lookup := (@lookup base_type_code Name var Context).
-
- Global Instance context_equiv_Equivalence : Equivalence context_equiv | 10.
- Proof. split; repeat intro; congruence. Qed.
- Global Instance context_equiv_Reflexive : Reflexive context_equiv | 10 := _.
- Global Instance context_equiv_Symmetric : Symmetric context_equiv | 10 := _.
- Global Instance context_equiv_Transitive : Transitive context_equiv | 10 := _.
-
- Local Ltac proper_t n n' t t' :=
- destruct (dec (n = n')), (dec (t = t')); subst;
- repeat first [ reflexivity
- | rewrite lookupb_extendb_same by assumption
- | rewrite lookupb_extendb_different by assumption
- | rewrite lookupb_extendb_wrong_type by assumption
- | rewrite lookupb_removeb_same by assumption
- | rewrite lookupb_removeb_different by assumption
- | solve [ auto ] ].
-
- Global Instance extendb_Proper : forall {t}, Proper (context_equiv ==> eq ==> eq ==> context_equiv) (@extendb t) | 10.
- Proof.
- intros t ??? n n0 ???? n' t'; subst n0; subst; proper_t n n' t t'.
- Qed.
- Global Instance removeb_Proper : forall {t}, Proper (context_equiv ==> eq ==> context_equiv) (@removeb t) | 10.
- Proof.
- intros t ??? n n0 ? n' t'; subst n0; subst; proper_t n n' t t'.
- Qed.
- Global Instance lookupb_Proper : forall {t}, Proper (context_equiv ==> eq ==> eq) (@lookupb t) | 10.
- Proof. repeat intro; subst; auto. Qed.
-
- Local Ltac proper_t2 t :=
- let IHt1 := fresh "IHt1" in
- let IHt2 := fresh "IHt2" in
- induction t as [ | | ? IHt1 ? IHt2];
- simpl;
- repeat match goal with
- | [ |- context[fun a b => ?f a b] ]
- => change (fun a b => f a b) with f
- end;
- [ try exact _
- | repeat intro; auto
- | repeat intro; subst;
- destruct_head_prod;
- first [ eapply IHt2; [ eapply IHt1 | .. ]; auto
- | idtac ] ].
-
- Global Instance extend_Proper : forall {t}, Proper (context_equiv ==> eq ==> eq ==> context_equiv) (@extend t) | 10.
- Proof. intro t; proper_t2 t. Qed.
- Global Instance remove_Proper : forall {t}, Proper (context_equiv ==> eq ==> context_equiv) (@remove t) | 10.
- Proof. intro t; proper_t2 t. Qed.
-
- Global Instance lookup_Proper : forall {t}, Proper (context_equiv ==> eq ==> eq) (@lookup t) | 10.
- Proof.
- intro t; proper_t2 t.
- repeat match goal with
- | [ |- context G[lookup (?A * ?B) ?ctx (?na, ?nb)] ]
- => let G' := context G[match lookup A ctx na, lookup B ctx nb with
- | Some a, Some b => Some (a, b)
- | _, _ => None
- end] in
- change G'
- end.
- break_innermost_match;
- repeat match goal with
- | _ => progress subst
- | _ => progress inversion_option
- | _ => reflexivity
- | [ IHt2 : Proper _ (lookup t2), H : lookup _ _ _ = ?x, H' : lookup _ _ _ = ?y |- _ ]
- => assert (x = y)
- by (rewrite <- H, <- H'; first [ eapply IHt1 | eapply IHt2 ]; (assumption || reflexivity));
- clear H H'
- end.
- Qed.
-End with_context.
-
-Section language.
- Context {base_type_code}
- {base_type_code_dec : DecidableRel (@eq base_type_code)}
- {Name}
- {Name_dec : DecidableRel (@eq Name)}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}
- {Context : @Context base_type_code Name interp_base_type}
- {ContextOk : ContextOk Context}.
-
- Local Notation context_equiv := (@context_equiv base_type_code Name interp_base_type Context).
- Local Notation interpf := (@interpf base_type_code interp_base_type op Name Context interp_op).
- Local Notation interp := (@interp base_type_code interp_base_type op Name Context interp_op).
-
- Global Instance interpf_Proper {t} : Proper (context_equiv ==> eq ==> eq) (@interpf t).
- Proof.
- intros c0 c1 Hc e e' ?; subst e'; revert c0 c1 Hc.
- induction e;
- repeat first [ progress subst
- | reflexivity
- | progress inversion_option
- | progress simpl in *
- | progress unfold LetIn.Let_In
- | intros; eapply lookupb_Proper; (assumption || reflexivity)
- | intros; eapply extendb_Proper; (assumption || reflexivity)
- | intros; eapply lookup_Proper; (assumption || reflexivity)
- | intros; eapply extend_Proper; (assumption || reflexivity)
- | intros; erewrite_hyp *; [ | eassumption ]
- | intros; erewrite_hyp *; [ reflexivity | ]
- | break_innermost_match_step
- | match goal with
- | [ H : interpf _ = ?x, H' : interpf _ = ?y |- _ ]
- => assert (x = y) by (rewrite <- H, <- H'; eauto); clear H H'
- end ].
- Qed.
-
- Global Instance interp_Proper {t} : Proper (context_equiv ==> eq ==> eq ==> eq) (@interp t).
- Proof.
- intros ??? e e' ????; subst e'; subst.
- eapply interpf_Proper; [ eapply extend_Proper; auto | reflexivity ].
- Qed.
-End language.
diff --git a/src/Compilers/Named/ContextProperties/SmartMap.v b/src/Compilers/Named/ContextProperties/SmartMap.v
deleted file mode 100644
index 2bfba38dc..000000000
--- a/src/Compilers/Named/ContextProperties/SmartMap.v
+++ /dev/null
@@ -1,299 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.Tactics.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Section with_context.
- Context {base_type_code Name var} (Context : @Context base_type_code Name var)
- (base_type_code_dec : DecidableRel (@eq base_type_code))
- (Name_dec : DecidableRel (@eq Name))
- (ContextOk : ContextOk Context).
-
- Local Notation find_Name := (@find_Name base_type_code Name Name_dec).
- Local Notation find_Name_and_val := (@find_Name_and_val base_type_code Name base_type_code_dec Name_dec).
-
- Hint Rewrite (@find_Name_and_val_default_to_None _ base_type_code_dec _ Name_dec) using congruence : ctx_db.
- Hint Rewrite (@find_Name_and_val_different _ base_type_code_dec _ Name_dec) using assumption : ctx_db.
- Hint Rewrite (@find_Name_and_val_wrong_type _ base_type_code_dec _ Name_dec) using congruence : ctx_db.
-
- Lemma find_Name_and_val_flatten_binding_list
- {var' var'' t n T N V1 V2 v1 v2}
- (H1 : @find_Name_and_val var' t n T N V1 None = Some v1)
- (H2 : @find_Name_and_val var'' t n T N V2 None = Some v2)
- : List.In (existT (fun t => (var' t * var'' t)%type) t (v1, v2)) (Wf.flatten_binding_list V1 V2).
- Proof using Type.
- induction T;
- [ | | specialize (IHT1 (fst N) (fst V1) (fst V2));
- specialize (IHT2 (snd N) (snd V1) (snd V2)) ];
- repeat first [ find_Name_and_val_default_to_None_step
- | rewrite List.in_app_iff
- | t_step ].
- Qed.
-
- Lemma find_Name_SmartFlatTypeMapInterp2_None_iff {var' n f T V N}
- : @find_Name n (SmartFlatTypeMap f (t:=T) V)
- (SmartFlatTypeMapInterp2 (var':=var') (fun _ _ (n' : Name) => n') V N) = None
- <-> find_Name n N = None.
- Proof using Type.
- split;
- (induction T;
- [ | | specialize (IHT1 (fst V) (fst N));
- specialize (IHT2 (snd V) (snd N)) ]);
- t.
- Qed.
- Hint Rewrite @find_Name_SmartFlatTypeMapInterp2_None_iff : ctx_db.
- Lemma find_Name_SmartFlatTypeMapInterp2_None {var' n f T V N}
- : @find_Name n (SmartFlatTypeMap f (t:=T) V)
- (SmartFlatTypeMapInterp2 (var':=var') (fun _ _ (n' : Name) => n') V N) = None
- -> find_Name n N = None.
- Proof using Type. apply find_Name_SmartFlatTypeMapInterp2_None_iff. Qed.
- Hint Rewrite @find_Name_SmartFlatTypeMapInterp2_None using eassumption : ctx_db.
- Lemma find_Name_SmartFlatTypeMapInterp2_None' {var' n f T V N}
- : find_Name n N = None
- -> @find_Name n (SmartFlatTypeMap f (t:=T) V)
- (SmartFlatTypeMapInterp2 (var':=var') (fun _ _ (n' : Name) => n') V N) = None.
- Proof using Type. apply find_Name_SmartFlatTypeMapInterp2_None_iff. Qed.
- Lemma find_Name_SmartFlatTypeMapInterp2_None_Some_wrong {var' n f T V N v}
- : @find_Name n (SmartFlatTypeMap f (t:=T) V)
- (SmartFlatTypeMapInterp2 (var':=var') (fun _ _ (n' : Name) => n') V N) = None
- -> find_Name n N = Some v
- -> False.
- Proof using Type.
- intro; erewrite find_Name_SmartFlatTypeMapInterp2_None by eassumption; congruence.
- Qed.
- Local Hint Resolve @find_Name_SmartFlatTypeMapInterp2_None_Some_wrong.
-
- Lemma find_Name_SmartFlatTypeMapInterp2 {var' n f T V N}
- : @find_Name n (SmartFlatTypeMap f (t:=T) V)
- (SmartFlatTypeMapInterp2 (var':=var') (fun _ _ (n' : Name) => n') V N)
- = match find_Name n N with
- | Some t' => match find_Name_and_val t' n N V None with
- | Some v => Some (f t' v)
- | None => None
- end
- | None => None
- end.
- Proof using Type.
- induction T;
- [ | | specialize (IHT1 (fst V) (fst N));
- specialize (IHT2 (snd V) (snd N)) ].
- { t. }
- { t. }
- { repeat first [ fin_t_step
- | inversion_step
- | rewrite_lookupb_extendb_step
- | misc_t_step
- | refolder_t_step ].
- repeat match goal with
- | [ |- context[match @find_Name ?n ?T ?N with _ => _ end] ]
- => destruct (@find_Name n T N) eqn:?
- | [ H : context[match @find_Name ?n ?T ?N with _ => _ end] |- _ ]
- => destruct (@find_Name n T N) eqn:?
- end;
- repeat first [ fin_t_step
- | rewriter_t_step
- | fin_t_late_step ]. }
- Qed.
-
- Lemma find_Name_and_val__SmartFlatTypeMapInterp2__SmartFlatTypeMapUnInterp__Some_Some_alt
- {var' var'' var''' t b n f g T V N X v v'}
- (Hfg
- : forall (V : var' t) (X : var'' (f t V)) (H : f t V = f t b),
- g t b (eq_rect (f t V) var'' X (f t b) H) = g t V X)
- : @find_Name_and_val
- var'' (f t b) n (SmartFlatTypeMap f (t:=T) V)
- (SmartFlatTypeMapInterp2 (var':=var') (fun _ _ (n' : Name) => n') V N) X None = Some v
- -> @find_Name_and_val
- var''' t n T N (SmartFlatTypeMapUnInterp (f:=f) g X) None = Some v'
- -> g t b v = v'.
- Proof using Type.
- induction T;
- [ | | specialize (IHT1 (fst V) (fst N) (fst X));
- specialize (IHT2 (snd V) (snd N) (snd X)) ];
- repeat first [ find_Name_and_val_default_to_None_step
- | t_step
- | match goal with
- | [ H : _ |- _ ]
- => progress rewrite find_Name_and_val_different in H
- by solve [ congruence
- | apply find_Name_SmartFlatTypeMapInterp2_None'; assumption ]
- end ].
- Qed.
-
- Lemma find_Name_and_val__SmartFlatTypeMapInterp2__SmartFlatTypeMapUnInterp__Some_Some
- {var' var'' var''' t b n f g T V N X v v'}
- : @find_Name_and_val
- var'' (f t b) n (SmartFlatTypeMap f (t:=T) V)
- (SmartFlatTypeMapInterp2 (var':=var') (fun _ _ (n' : Name) => n') V N) X None = Some v
- -> @find_Name_and_val
- _ t n T N V None = Some b
- -> @find_Name_and_val
- var''' t n T N (SmartFlatTypeMapUnInterp (f:=f) g X) None = Some v'
- -> g t b v = v'.
- Proof using Type.
- induction T;
- [ | | specialize (IHT1 (fst V) (fst N) (fst X));
- specialize (IHT2 (snd V) (snd N) (snd X)) ];
- repeat first [ find_Name_and_val_default_to_None_step
- | t_step
- | match goal with
- | [ H : _ |- _ ]
- => progress rewrite find_Name_and_val_different in H
- by solve [ congruence
- | apply find_Name_SmartFlatTypeMapInterp2_None'; assumption ]
- end ].
- Qed.
-
- Lemma interp_flat_type_rel_pointwise__find_Name_and_val
- {var' var'' t n T N x y R v0 v1}
- (H0 : @find_Name_and_val var' t n T N x None = Some v0)
- (H1 : @find_Name_and_val var'' t n T N y None = Some v1)
- (HR : interp_flat_type_rel_pointwise R x y)
- : R _ v0 v1.
- Proof using Type.
- induction T;
- [ | | specialize (IHT1 (fst N) (fst x) (fst y));
- specialize (IHT2 (snd N) (snd x) (snd y)) ];
- repeat first [ find_Name_and_val_default_to_None_step
- | t_step ].
- Qed.
-
- Lemma find_Name_and_val_SmartFlatTypeMapUnInterp2_Some_Some
- {var' var'' var''' f g}
- {T}
- {N : interp_flat_type (fun _ : base_type_code => Name) T}
- {B : interp_flat_type var' T}
- {V : interp_flat_type var'' (SmartFlatTypeMap (t:=T) f B)}
- {n : Name}
- {t : base_type_code}
- {v : var''' t}
- {b}
- {h} {i : forall v, var'' (f _ (h v))}
- (Hn : find_Name n N = Some t)
- (Hf : find_Name_and_val t n N (SmartFlatTypeMapUnInterp2 g V) None = Some v)
- (Hb : find_Name_and_val t n N B None = Some b)
- (Hig : forall B V,
- existT _ (h (g _ B V)) (i (g _ B V))
- = existT _ B V
- :> { b : _ & var'' (f _ b)})
- (N' := SmartFlatTypeMapInterp2 (var'':=fun _ => Name) (f:=f) (fun _ _ n => n) _ N)
- : b = h v /\ find_Name_and_val (f t (h v)) n N' V None = Some (i v).
- Proof using Type.
- induction T;
- [ | | specialize (IHT1 (fst N) (fst B) (fst V));
- specialize (IHT2 (snd N) (snd B) (snd V)) ];
- repeat first [ find_Name_and_val_default_to_None_step
- | lazymatch goal with
- | [ H : context[find_Name ?n (@SmartFlatTypeMapInterp2 _ ?var' _ _ ?f _ ?T ?V ?N)] |- _ ]
- => setoid_rewrite find_Name_SmartFlatTypeMapInterp2 in H
- end
- | t_step
- | match goal with
- | [ Hhg : forall B V, existT _ (?h (?g ?t B V)) _ = existT _ B _ |- context[?h (?g ?t ?B' ?V')] ]
- => specialize (Hhg B' V'); generalize dependent (g t B' V')
- end ].
- Qed.
-End with_context.
-
-Section with_context2.
- Context {base_type_code1 base_type_code2 Name var1 var2}
- {Context1 : @Context base_type_code1 Name var1}
- {Context2 : @Context base_type_code2 Name var2}
- {base_type_code1_dec : DecidableRel (@eq base_type_code1)}
- {base_type_code2_dec : DecidableRel (@eq base_type_code2)}
- {Name_dec : DecidableRel (@eq Name)}
- {Context1Ok : ContextOk Context1}
- {Context2Ok : ContextOk Context2}
- {f_base : base_type_code1 -> base_type_code2}
- {cast_backb: forall t, var2 (f_base t) -> var1 t}.
-
- Let cast_back : forall t, interp_flat_type var2 (lift_flat_type _ t) -> interp_flat_type var1 t
- := fun t => untransfer_interp_flat_type f_base cast_backb.
-
- Local Notation find_Name1 := (@find_Name base_type_code1 Name Name_dec).
- Local Notation find_Name_and_val1 := (@find_Name_and_val base_type_code1 Name base_type_code1_dec Name_dec).
- Local Notation find_Name2 := (@find_Name base_type_code2 Name Name_dec).
- Local Notation find_Name_and_val2 := (@find_Name_and_val base_type_code2 Name base_type_code2_dec Name_dec).
-
- Lemma find_Name_transfer_interp_flat_type {T n N}
- : find_Name2
- n
- (transfer_interp_flat_type
- f_base (fun _ (n : Name) => n)
- N)
- = option_map f_base (find_Name1 (T:=T) n N).
- Proof.
- induction T; simpl in *;
- [ | | specialize (IHT1 (fst N));
- specialize (IHT2 (snd N)) ];
- repeat first [ reflexivity
- | misc_t_step
- | rewriter_t_step
- | progress cbv [option_map] in *
- | fin_t_late_step
- | break_t_step ].
- Qed.
-
- Local Ltac t' :=
- repeat first [ reflexivity
- | progress subst
- | progress inversion_step
- | progress intros
- | progress simpl in *
- | rewrite cast_if_eq_refl
- | break_innermost_match_step
- | break_innermost_match_hyps_step
- | solve [ auto ]
- | specializer_t_step
- | symmetry; find_Name_and_val_default_to_None_step; symmetry;
- rewrite find_Name_transfer_interp_flat_type
- | find_Name_and_val_default_to_None_step;
- match goal with
- | [ H : ?x = None |- context[?x] ] => rewrite H
- end
- | progress cbv [option_map] in *
- | congruence ].
-
- Lemma find_Name_and_val_transfer_interp_flat_type {t T} {n : Name} {N' N} {default default'}
- (H' : find_Name Name_dec n N = Some t)
- : find_Name_and_val1 t n N (cast_back T N') default
- = option_map
- (cast_backb _)
- (find_Name_and_val2
- (f_base t) n
- (transfer_interp_flat_type
- f_base
- (fun _ (n : Name) => n) N)
- N'
- default').
- Proof.
- revert default default'; induction T; intros;
- [ | | specialize (IHT1 (fst N') (fst N));
- specialize (IHT2 (snd N') (snd N)) ];
- t'.
- Qed.
-
- Lemma find_Name_and_val_transfer_interp_flat_type_None {t T} {n : Name} {N' N} {default'}
- (H' : find_Name Name_dec n N = None)
- : find_Name_and_val2
- (var':=var2)
- (f_base t) n
- (transfer_interp_flat_type
- (t:=T)
- f_base
- (fun _ (n : Name) => n) N)
- N'
- default'
- = default'.
- Proof.
- revert default'; induction T; intros;
- [ | | specialize (IHT1 (fst N') (fst N));
- specialize (IHT2 (snd N') (snd N)) ];
- t'.
- Qed.
-End with_context2.
diff --git a/src/Compilers/Named/ContextProperties/Tactics.v b/src/Compilers/Named/ContextProperties/Tactics.v
deleted file mode 100644
index b7e11a769..000000000
--- a/src/Compilers/Named/ContextProperties/Tactics.v
+++ /dev/null
@@ -1,101 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.HProp.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-
-Ltac fin_t_step :=
- solve [ reflexivity ].
-Ltac fin_t_late_step :=
- solve [ tauto
- | congruence
- | eauto
- | exfalso; unfold not in *; eauto ].
-Ltac inversion_step :=
- first [ progress subst
- | match goal with
- | [ H := _ |- _ ] => subst H
- | [ H : ?x = ?y |- _ ] => subst x || subst y
- end
- | progress inversion_option
- | progress inversion_sigma
- | progress inversion_prod
- | progress destruct_head' and
- | progress eliminate_hprop_eq
- | match goal with
- | [ H : ?T, H' : ?T |- _ ] => clear H'
- | [ H : ?x = ?x |- _ ] => clear H
- end
- | progress split_and ].
-Ltac rewrite_lookupb_extendb_step :=
- first [ rewrite lookupb_extendb_different by congruence
- | rewrite lookupb_extendb_same
- | rewrite lookupb_removeb_same
- | rewrite lookupb_extendb_wrong_type by assumption
- | rewrite lookupb_removeb_different by congruence ].
-Ltac specializer_t_step :=
- match goal with
- | _ => progress specialize_by_assumption
- | [ H : ?x = ?x -> _ |- _ ] => specialize (H eq_refl)
- | [ H : ?T, H' : ?T |- _ ] => clear H'
- | [ H : forall v, Some _ = Some v -> _ |- _ ]
- => specialize (H _ eq_refl)
- | [ H : forall v, Some v = Some _ -> _ |- _ ]
- => specialize (H _ eq_refl)
- end.
-Ltac misc_t_step :=
- first [ progress intros
- | progress simpl in * ].
-Ltac break_t_step :=
- first [ progress break_innermost_match_step
- | progress unfold cast_if_eq in *
- | match goal with
- | [ H : context[match _ with _ => _ end] |- _ ]
- => revert H; progress break_innermost_match_step
- | [ |- _ /\ _ ] => split
- | [ |- _ <-> _ ] => split
- | [ |- ~ _ ] => intro
- end
- | progress destruct_head' ex
- | progress destruct_head' or ].
-Ltac refolder_t_step :=
- let myfold_in_star c :=
- (let c' := (eval cbv [interp_flat_type] in c) in
- change c' with c in * ) in
- first [ match goal with
- | [ var : ?base_type_code -> Type |- _ ]
- => progress myfold_in_star (@interp_flat_type base_type_code var)
- | [ base_type_code : Type, Name : Type |- _ ]
- => progress myfold_in_star (@interp_flat_type base_type_code (fun _ => Name))
-
- end ].
-Ltac rewriter_t_step :=
- first [ match goal with
- | [ H : _ |- _ ] => rewrite H by (assumption || congruence)
- | [ H : _ |- _ ] => etransitivity; [ | rewrite H by (assumption || congruence); reflexivity ]
- | [ H : ?x = Some _, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : ?x = None, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : ?x = ?a :> option _, H' : ?x = ?b :> option _ |- _ ]
- => assert (a = b)
- by (transitivity x; [ symmetry | ]; assumption);
- clear H'
- | _ => progress autorewrite with ctx_db in *
- end ].
-Ltac t_step :=
- first [ fin_t_step
- | inversion_step
- | rewrite_lookupb_extendb_step
- | specializer_t_step
- | misc_t_step
- | break_t_step
- | refolder_t_step
- | rewriter_t_step
- | fin_t_late_step ].
-Ltac t := repeat t_step.
diff --git a/src/Compilers/Named/CountLets.v b/src/Compilers/Named/CountLets.v
deleted file mode 100644
index 1daa8a286..000000000
--- a/src/Compilers/Named/CountLets.v
+++ /dev/null
@@ -1,49 +0,0 @@
-(** * Counts how many binders there are *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.CountLets.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
-
- Local Notation count_pairs := (@count_pairs base_type_code).
-
- Local Notation exprf := (@Named.exprf base_type_code op Name).
- Local Notation expr := (@Named.expr base_type_code op Name).
-
- Section gen.
- Context (count_type_let : flat_type -> nat).
- Context (count_type_abs : flat_type -> nat).
-
- Fixpoint count_lets_genf {t} (e : exprf t) : nat
- := match e with
- | LetIn tx _ _ _ eC
- => count_type_let tx + @count_lets_genf _ eC
- | Op _ _ _ e => @count_lets_genf _ e
- | Pair _ ex _ ey => @count_lets_genf _ ex + @count_lets_genf _ ey
- | _ => 0
- end.
- Definition count_lets_gen {t} (e : expr t) : nat
- := match e with
- | Abs tx _ _ f => count_type_abs tx + @count_lets_genf _ f
- end.
- End gen.
-
- Definition count_let_bindersf {t} (e : exprf t) : nat
- := count_lets_genf count_pairs e.
- Definition count_letsf {t} (e : exprf t) : nat
- := count_lets_genf (fun _ => 1) e.
- Definition count_let_binders {t} (e : expr t) : nat
- := count_lets_gen count_pairs (fun _ => 0) e.
- Definition count_lets {t} (e : expr t) : nat
- := count_lets_gen (fun _ => 1) (fun _ => 0) e.
- Definition count_binders {t} (e : expr t) : nat
- := count_lets_gen count_pairs count_pairs e.
-End language.
diff --git a/src/Compilers/Named/DeadCodeElimination.v b/src/Compilers/Named/DeadCodeElimination.v
deleted file mode 100644
index 340c5dd0d..000000000
--- a/src/Compilers/Named/DeadCodeElimination.v
+++ /dev/null
@@ -1,54 +0,0 @@
-(** * PHOAS → Named Representation of Gallina *)
-Require Import Coq.PArith.BinPos Coq.Lists.List.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Compilers.Named.RegisterAssign.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.EstablishLiveness.
-Require Import Crypto.Compilers.CountLets.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.LetIn.
-
-Local Notation eta x := (fst x, snd x).
-
-Local Open Scope ctype_scope.
-Local Open Scope nexpr_scope.
-Local Open Scope expr_scope.
-Section language.
- Context (base_type_code : Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (Name : Type)
- {base_type_code_beq : base_type_code -> base_type_code -> bool}
- (base_type_code_bl : forall t1 t2, base_type_code_beq t1 t2 = true -> t1 = t2)
- {Context : Context Name (fun _ : base_type_code => positive)}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op (fun _ => Name)).
- Local Notation expr := (@expr base_type_code op (fun _ => Name)).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation nexprf := (@Named.exprf base_type_code op Name).
- Local Notation nexpr := (@Named.expr base_type_code op Name).
-
- Local Notation PContext var := (@PositiveContext base_type_code var base_type_code_beq base_type_code_bl).
-
- Definition EliminateDeadCode
- {t} (e : @Named.expr base_type_code op _ t) (ls : list Name)
- : option (nexpr t)
- := Let_In (insert_dead_names (Context:=PositiveContext_nd) None e ls) (* help vm_compute by factoring this out *)
- (fun names => register_reassign (InContext:=PContext _) (ReverseContext:=Context) Pos.eqb empty empty e names).
-
- Definition CompileAndEliminateDeadCode
- {t} (e : Expr t) (ls : list Name)
- : option (nexpr t)
- := let e := compile (Name:=positive) (e _) (List.map Pos.of_nat (seq 1 (CountBinders e))) in
- match e with
- | Some e => EliminateDeadCode e ls
- | None => None
- end.
-End language.
-
-Global Arguments EliminateDeadCode {_ _ _ _ _ _ t} e ls.
-Global Arguments CompileAndEliminateDeadCode {_ _ _ _ _ _ t} e ls.
diff --git a/src/Compilers/Named/DeadCodeEliminationInterp.v b/src/Compilers/Named/DeadCodeEliminationInterp.v
deleted file mode 100644
index f822e6e29..000000000
--- a/src/Compilers/Named/DeadCodeEliminationInterp.v
+++ /dev/null
@@ -1,67 +0,0 @@
-Require Import Coq.PArith.BinPos.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.RegisterAssignInterp.
-Require Import Crypto.Compilers.Named.DeadCodeElimination.
-Require Import Crypto.Util.Decidable.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- {interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {Name : Type}
- {InContext : Context Name (fun _ : base_type_code => BinNums.positive)}
- {InContextOk : ContextOk InContext}
- {Name_beq : Name -> Name -> bool}
- {InterpContext : Context Name interp_base_type}
- {InterpContextOk : ContextOk InterpContext}
- {base_type_code_beq : base_type_code -> base_type_code -> bool}
- (base_type_code_bl : forall t1 t2, base_type_code_beq t1 t2 = true -> t1 = t2)
- (base_type_code_lb : forall t1 t2, t1 = t2 -> base_type_code_beq t1 t2 = true)
- (Name_bl : forall n1 n2, Name_beq n1 n2 = true -> n1 = n2)
- (Name_lb : forall n1 n2, n1 = n2 -> Name_beq n1 n2 = true).
-
- Local Notation EliminateDeadCode := (@EliminateDeadCode base_type_code op Name _ base_type_code_bl InContext).
- Local Notation PContext var := (@PositiveContext base_type_code var base_type_code_beq base_type_code_bl).
-
- Local Instance Name_dec : DecidableRel (@eq Name)
- := dec_rel_of_bool_dec_rel Name_beq Name_bl Name_lb.
- Local Instance base_type_code_dec : DecidableRel (@eq base_type_code)
- := dec_rel_of_bool_dec_rel base_type_code_beq base_type_code_bl base_type_code_lb.
-
- Lemma interp_EliminateDeadCode
- t e new_names
- (ctxi_interp : PContext _)
- (ctxr_interp : InterpContext)
- eout v1 v2 x
- : @EliminateDeadCode t e new_names = Some eout
- -> interp (interp_op:=interp_op) (ctx:=ctxr_interp) eout x = Some v1
- -> interp (interp_op:=interp_op) (ctx:=ctxi_interp) e x = Some v2
- -> v1 = v2.
- Proof using InContextOk InterpContextOk Name_bl Name_lb base_type_code_lb.
- eapply @interp_register_reassign;
- first [ assumption
- | apply @PositiveContextOk; assumption
- | apply Peqb_true_eq
- | apply Pos.eqb_eq
- | exact _
- | intros *; rewrite !lookupb_empty by (try apply @PositiveContextOk; assumption);
- congruence ].
- Qed.
-
- Lemma InterpEliminateDeadCode
- t e new_names
- eout
- v1 v2 x
- : @EliminateDeadCode t e new_names = Some eout
- -> Interp (Context:=InterpContext) (interp_op:=interp_op) eout x = Some v1
- -> Interp (Context:=PContext _) (interp_op:=interp_op) e x = Some v2
- -> v1 = v2.
- Proof using InContextOk InterpContextOk Name_bl Name_lb base_type_code_lb.
- apply interp_EliminateDeadCode.
- Qed.
-End language.
diff --git a/src/Compilers/Named/EstablishLiveness.v b/src/Compilers/Named/EstablishLiveness.v
deleted file mode 100644
index 706327918..000000000
--- a/src/Compilers/Named/EstablishLiveness.v
+++ /dev/null
@@ -1,105 +0,0 @@
-(** * Compute a list of liveness values for each binding *)
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.CountLets.
-Require Import Crypto.Util.ListUtil.
-
-Local Notation eta x := (fst x, snd x).
-
-Local Open Scope ctype_scope.
-Delimit Scope nexpr_scope with nexpr.
-
-Inductive liveness := live | dead.
-Fixpoint merge_liveness (ls1 ls2 : list liveness) :=
- match ls1, ls2 with
- | cons x xs, cons y ys
- => cons match x, y with
- | live, _
- | _, live
- => live
- | dead, dead
- => dead
- end
- (@merge_liveness xs ys)
- | nil, ls
- | ls, nil
- => ls
- end.
-
-Section language.
- Context (base_type_code : Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
-
- Section internal.
- Context (Name : Type)
- (OutName : Type)
- {Context : Context Name (fun _ : base_type_code => list liveness)}.
-
- Definition compute_livenessf_step
- (compute_livenessf : forall (ctx : Context) {t} (e : exprf Name t) (prefix : list liveness), list liveness)
- (ctx : Context)
- {t} (e : exprf Name t) (prefix : list liveness)
- : list liveness
- := match e with
- | TT => prefix
- | Var t' name => match lookup (Tbase t') ctx name with
- | Some ls => ls
- | _ => nil
- end
- | Op _ _ op args
- => @compute_livenessf ctx _ args prefix
- | LetIn tx n ex _ eC
- => let lx := @compute_livenessf ctx _ ex prefix in
- let lx := merge_liveness lx (prefix ++ repeat live (count_pairs tx)) in
- let ctx := extend ctx n (SmartValf _ (fun _ => lx) tx) in
- @compute_livenessf ctx _ eC (prefix ++ repeat dead (count_pairs tx))
- | Pair _ ex _ ey
- => merge_liveness (@compute_livenessf ctx _ ex prefix)
- (@compute_livenessf ctx _ ey prefix)
- end.
-
- Fixpoint compute_livenessf ctx {t} e prefix
- := @compute_livenessf_step (@compute_livenessf) ctx t e prefix.
-
- Definition compute_liveness (ctx : Context)
- {t} (e : expr Name t) (prefix : list liveness)
- : list liveness
- := match e with
- | Abs src _ n f
- => let prefix := prefix ++ repeat live (count_pairs src) in
- let ctx := extend (t:=src) ctx n (SmartValf _ (fun _ => prefix) src) in
- compute_livenessf ctx f prefix
- end.
-
- Section insert_dead.
- Context (default_out : option OutName).
-
- Fixpoint insert_dead_names_gen (ls : list liveness) (lsn : list OutName)
- : list (option OutName)
- := match ls with
- | nil => nil
- | cons live xs
- => match lsn with
- | cons n lsn' => Some n :: @insert_dead_names_gen xs lsn'
- | nil => default_out :: @insert_dead_names_gen xs nil
- end
- | cons dead xs
- => None :: @insert_dead_names_gen xs lsn
- end.
- Definition insert_dead_names {t} (e : expr Name t)
- := insert_dead_names_gen (compute_liveness empty e nil).
- End insert_dead.
- End internal.
-End language.
-
-Global Arguments compute_livenessf {_ _ _ _} ctx {t} e prefix.
-Global Arguments compute_liveness {_ _ _ _} ctx {t} e prefix.
-Global Arguments insert_dead_names {_ _ _ _ _} default_out {t} e lsn.
diff --git a/src/Compilers/Named/ExprInversion.v b/src/Compilers/Named/ExprInversion.v
deleted file mode 100644
index 783341f26..000000000
--- a/src/Compilers/Named/ExprInversion.v
+++ /dev/null
@@ -1,21 +0,0 @@
-Require Import Crypto.Compilers.TypeInversion.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-
-Module Export Named.
- Ltac invert_one_expr e :=
- preinvert_one_type e;
- intros ? e;
- destruct e;
- try exact I.
-
- Ltac invert_expr_step :=
- match goal with
- | [ e : Named.exprf _ _ _ (Tbase _) |- _ ] => invert_one_expr e
- | [ e : Named.exprf _ _ _ (Prod _ _) |- _ ] => invert_one_expr e
- | [ e : Named.exprf _ _ _ Unit |- _ ] => invert_one_expr e
- | [ e : Named.expr _ _ _ (Arrow _ _) |- _ ] => invert_one_expr e
- end.
-
- Ltac invert_expr := repeat invert_expr_step.
-End Named.
diff --git a/src/Compilers/Named/FMapContext.v b/src/Compilers/Named/FMapContext.v
deleted file mode 100644
index b69d78d09..000000000
--- a/src/Compilers/Named/FMapContext.v
+++ /dev/null
@@ -1,70 +0,0 @@
-Require Import Coq.Bool.Sumbool.
-Require Import Coq.FSets.FMapInterface.
-Require Import Coq.FSets.FMapFacts.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Equality.
-
-Module FMapContextFun (E : DecidableType) (W : WSfun E).
- Module Import Properties := WProperties_fun E W.
- Import F.
- Section ctx.
- Context (E_eq_l : forall x y, E.eq x y -> x = y)
- base_type_code (var : base_type_code -> Type)
- (base_type_code_beq : base_type_code -> base_type_code -> bool)
- (base_type_code_bl_transparent : forall x y, base_type_code_beq x y = true -> x = y)
- (base_type_code_lb : forall x y, x = y -> base_type_code_beq x y = true).
-
- Definition var_cast {a b} (x : var a) : option (var b)
- := match Sumbool.sumbool_of_bool (base_type_code_beq a b), Sumbool.sumbool_of_bool (base_type_code_beq b b) with
- | left pf, left pf' => match eq_trans (base_type_code_bl_transparent _ _ pf) (eq_sym (base_type_code_bl_transparent _ _ pf')) with
- | eq_refl => Some x
- end
- | right _, _ | _, right _ => None
- end.
- Definition FMapContext : @Context base_type_code W.key var
- := {| ContextT := W.t { t : _ & var t };
- lookupb t ctx n
- := match W.find n ctx with
- | Some (existT t' v)
- => var_cast v
- | None => None
- end;
- extendb t ctx n v
- := W.add n (existT _ t v) ctx;
- removeb t ctx n
- := W.remove n ctx;
- empty := W.empty _ |}.
- Lemma FMapContextOk : @ContextOk base_type_code W.key var FMapContext.
- Proof using E_eq_l base_type_code_lb.
- split;
- repeat first [ reflexivity
- | progress simpl in *
- | progress intros
- | rewrite add_eq_o by reflexivity
- | progress rewrite ?add_neq_o, ?remove_neq_o, ?remove_eq_o in * by intuition
- | progress rewrite empty_o in *
- | progress unfold var_cast
- | progress break_innermost_match_step
- | rewrite concat_pV
- | congruence
- | rewrite base_type_code_lb in * by reflexivity
- | break_innermost_match_hyps_step
- | progress unfold var_cast in * ].
- Qed.
- End ctx.
-
- Section ctx_nd.
- Context {base_type_code var : Type}.
-
- Definition FMapContext_nd : @Context base_type_code W.key (fun _ => var)
- := {| ContextT := W.t var;
- lookupb t ctx n := W.find n ctx;
- extendb t ctx n v := W.add n v ctx;
- removeb t ctx n := W.remove n ctx;
- empty := W.empty _ |}.
- End ctx_nd.
-End FMapContextFun.
-
-Module FMapContext (W : WS) := FMapContextFun W.E W.
diff --git a/src/Compilers/Named/GetNames.v b/src/Compilers/Named/GetNames.v
deleted file mode 100644
index ce26b1bca..000000000
--- a/src/Compilers/Named/GetNames.v
+++ /dev/null
@@ -1,38 +0,0 @@
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-
-Local Open Scope list_scope.
-
-Section language.
- Context {base_type_code}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}.
-
- Local Notation exprf := (@exprf base_type_code op Name).
- Local Notation expr := (@expr base_type_code op Name).
-
- Fixpoint get_names_of_type {t} : interp_flat_type (fun _ : base_type_code => Name) t -> list Name
- := match t with
- | Tbase T => fun n => n::nil
- | Unit => fun _ => nil
- | Prod A B => fun ab : interp_flat_type _ A * interp_flat_type _ B
- => @get_names_of_type _ (fst ab) ++ @get_names_of_type _ (snd ab)
- end.
-
- Fixpoint get_namesf {t} (e : exprf t) : list Name
- := match e with
- | TT => nil
- | Var _ x => x::nil
- | Op _ _ opc args => @get_namesf _ args
- | LetIn tx nx ex tC eC
- => get_names_of_type nx ++ (* @get_namesf _ ex ++ *) @get_namesf _ eC
- | Pair _ ex _ ey
- => @get_namesf _ ex ++ @get_namesf _ ey
- end.
-
- Fixpoint get_names {t} (e : expr t) : list Name
- := match e with
- | Abs _ _ n f => get_names_of_type n ++ get_namesf f
- end.
-End language.
diff --git a/src/Compilers/Named/IdContext.v b/src/Compilers/Named/IdContext.v
deleted file mode 100644
index d82740cb4..000000000
--- a/src/Compilers/Named/IdContext.v
+++ /dev/null
@@ -1,27 +0,0 @@
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-
-Section language.
- Context {base_type_code Name}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- (Context : @Context base_type_code Name (fun _ => Name)).
-
- Fixpoint collect_binders {t} (e : Named.exprf base_type_code op Name t)
- : list { t : flat_type base_type_code & interp_flat_type (fun _ => Name) t }
- := match e with
- | TT => nil
- | Var t n => (existT _ (Tbase t) n) :: nil
- | Op t1 tR opc args => @collect_binders _ args
- | LetIn tx n ex tC eC
- => (existT _ tx n) :: @collect_binders tx ex ++ @collect_binders tC eC
- | Pair tx ex ty ey
- => @collect_binders tx ex ++ @collect_binders ty ey
- end%list.
- Definition idcontext {t} (e : Named.exprf base_type_code op Name t) : Context
- := List.fold_right
- (fun v ctx => extend ctx (projT2 v) (projT2 v))
- empty
- (collect_binders e).
-End language.
diff --git a/src/Compilers/Named/InterpSideConditions.v b/src/Compilers/Named/InterpSideConditions.v
deleted file mode 100644
index bd90eac23..000000000
--- a/src/Compilers/Named/InterpSideConditions.v
+++ /dev/null
@@ -1,54 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Notations.
-
-Module Export Named.
- Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {interp_base_type : base_type_code -> Type}
- {Context : Context Name interp_base_type}
- (interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d)
- (interped_op_side_conditions : forall s d, op s d -> interp_flat_type interp_base_type s -> pointed_Prop).
-
- Local Notation exprf := (@exprf base_type_code op Name).
- Local Notation expr := (@expr base_type_code op Name).
-
- Fixpoint interpf_side_conditions_gen {t} (ctx : Context) (e : exprf t)
- : option (pointed_Prop * interp_flat_type interp_base_type t)
- := match e with
- | TT => Some (trivial, tt)
- | Var t' x => option_map (fun v => (trivial, v)) (lookupb t' ctx x)
- | Op t1 tR opc args
- => match @interpf_side_conditions_gen _ ctx args with
- | Some (args_cond, argsv)
- => Some (args_cond /\ interped_op_side_conditions _ _ opc argsv, interp_op _ _ opc argsv)
- | None => None
- end
- | LetIn _ n ex _ eC
- => match @interpf_side_conditions_gen _ ctx ex with
- | Some (x_cond, x)
- => match @interpf_side_conditions_gen _ (extend ctx n x) eC with
- | Some (c_cond, cv)
- => Some (x_cond /\ c_cond, cv)
- | None => None
- end
- | None => None
- end
- | Pair _ ex _ ey
- => match @interpf_side_conditions_gen _ ctx ex, @interpf_side_conditions_gen _ ctx ey with
- | Some (x_cond, xv), Some (y_cond, yv) => Some (x_cond /\ y_cond, (xv, yv))
- | None, _ | _, None => None
- end
- end%pointed_prop.
- Definition interpf_side_conditions {t} ctx e : option pointed_Prop
- := option_map (@fst _ _) (@interpf_side_conditions_gen t ctx e).
- Definition interp_side_conditions {t} ctx (e : expr t) : interp_flat_type interp_base_type (domain t) -> option pointed_Prop
- := fun x => interpf_side_conditions (extend ctx (Abs_name e) x) (invert_Abs e).
- Definition InterpSideConditions {t} (e : expr t) : interp_flat_type interp_base_type (domain t) -> option pointed_Prop
- := interp_side_conditions empty e.
- End language.
-End Named.
diff --git a/src/Compilers/Named/InterpSideConditionsInterp.v b/src/Compilers/Named/InterpSideConditionsInterp.v
deleted file mode 100644
index 72bb9c9e0..000000000
--- a/src/Compilers/Named/InterpSideConditionsInterp.v
+++ /dev/null
@@ -1,45 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.InterpSideConditions.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Option.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {interp_base_type : base_type_code -> Type}
- {Context : Context Name interp_base_type}
- {interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {interped_op_side_conditions : forall s d, op s d -> interp_flat_type interp_base_type s -> pointed_Prop}.
-
- Local Notation exprf := (@exprf base_type_code op Name).
- Local Notation expr := (@expr base_type_code op Name).
-
- Lemma snd_interpf_side_conditions_gen_eq {t} {ctx : Context} {e}
- : interpf (t:=t) (ctx:=ctx) (interp_op:=interp_op) e
- = option_map (@snd _ _) (interpf_side_conditions_gen interp_op interped_op_side_conditions ctx e).
- Proof.
- revert ctx; induction e; intros;
- repeat first [ reflexivity
- | progress subst
- | progress inversion_option
- | progress unfold option_map, LetIn.Let_In in *
- | break_innermost_match_step
- | break_innermost_match_hyps_step
- | progress simpl in *
- | match goal with
- | [ H : forall ctx, interpf ?e = _, H' : context[interpf ?e] |- _ ]
- => rewrite H in H'
- | [ H : forall ctx, interpf ?e = _ |- context[interpf ?e] ]
- => rewrite H
- end ].
- Qed.
- Lemma snd_interpf_side_conditions_gen_Some {t} {ctx : Context} {e v}
- : interpf (t:=t) (ctx:=ctx) (interp_op:=interp_op) e = Some v
- <-> option_map (@snd _ _) (interpf_side_conditions_gen interp_op interped_op_side_conditions ctx e) = Some v.
- Proof. rewrite snd_interpf_side_conditions_gen_eq; reflexivity. Qed.
-End language.
diff --git a/src/Compilers/Named/InterpretToPHOAS.v b/src/Compilers/Named/InterpretToPHOAS.v
deleted file mode 100644
index 4e74fc59a..000000000
--- a/src/Compilers/Named/InterpretToPHOAS.v
+++ /dev/null
@@ -1,65 +0,0 @@
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Util.PointedProp.
-
-Local Notation eta_and x := (conj (let (a, b) := x in a) (let (a, b) := x in b)).
-
-Module Export Named.
- Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}.
- Section with_var.
- Context {var : base_type_code -> Type}
- {Context : Context Name var}
- (failb : forall t, @Syntax.exprf base_type_code op var (Tbase t)).
-
- Local Notation failf t (* : @Syntax.exprf base_type_code op var t*)
- := (SmartPairf (SmartValf _ failb t)).
-
- Fixpoint interpf_to_phoas
- (ctx : Context)
- {t} (e : @Named.exprf base_type_code op Name t)
- {struct e}
- : @Syntax.exprf base_type_code op var t
- := match e in Named.exprf _ _ _ t return @Syntax.exprf base_type_code op var t with
- | Named.Var t' x
- => match lookupb t' ctx x with
- | Some v => Var v
- | None => failf _
- end
- | Named.TT => TT
- | Named.Pair tx ex ty ey
- => Pair (@interpf_to_phoas ctx tx ex) (@interpf_to_phoas ctx ty ey)
- | Named.Op _ _ opc args
- => Op opc (@interpf_to_phoas ctx _ args)
- | Named.LetIn _ n ex _ eC
- => LetIn (@interpf_to_phoas ctx _ ex)
- (fun v
- => @interpf_to_phoas (extend ctx n v) _ eC)
- end.
-
- Definition interp_to_phoas
- (ctx : Context)
- {t} (e : @Named.expr base_type_code op Name t)
- : @Syntax.expr base_type_code op var (domain t -> codomain t)
- := Abs (fun v => interpf_to_phoas (extend ctx (Abs_name e) v) (invert_Abs e)).
- End with_var.
-
- Section all.
- Context {Context : forall var, @Context base_type_code Name var}
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t)).
- Definition InterpToPHOAS_gen
- (ctx : forall var, Context var)
- {t} (e : @Named.expr base_type_code op Name t)
- : @Syntax.Expr base_type_code op (domain t -> codomain t)
- := fun var => interp_to_phoas (failb var) (ctx var) e.
-
- Definition InterpToPHOAS {t} e
- := @InterpToPHOAS_gen (fun var => empty) t e.
- End all.
- End language.
-End Named.
diff --git a/src/Compilers/Named/InterpretToPHOASInterp.v b/src/Compilers/Named/InterpretToPHOASInterp.v
deleted file mode 100644
index 7fe50c994..000000000
--- a/src/Compilers/Named/InterpretToPHOASInterp.v
+++ /dev/null
@@ -1,89 +0,0 @@
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.InterpretToPHOAS.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {base_type_code_dec : DecidableRel (@eq base_type_code)}
- {Name_dec : DecidableRel (@eq Name)}
- {interp_base_type : base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}.
- Section with_context.
- Context {Context : Context Name interp_base_type}
- {ContextOk : ContextOk Context}
- (failb : forall t, @Syntax.exprf base_type_code op interp_base_type (Tbase t)).
-
- Lemma interpf_interpf_to_phoas
- (ctx : Context)
- {t} (e : @Named.exprf base_type_code op Name t)
- (Hwf : prop_of_option (Named.wff ctx e))
- : Named.interpf (interp_op:=interp_op) (ctx:=ctx) e
- = Some (Syntax.interpf interp_op (interpf_to_phoas failb ctx e)).
- Proof using Type.
- revert dependent ctx; induction e;
- repeat first [ progress intros
- | progress subst
- | progress inversion_option
- | progress destruct_head' and
- | progress break_innermost_match_step
- | progress unfold option_map, LetIn.Let_In in *
- | apply (f_equal (@Some _))
- | apply (f_equal (@interp_op _ _ _))
- | progress simpl in *
- | progress autorewrite with push_prop_of_option in *
- | solve [ eauto | congruence | tauto ]
- | match goal with
- | [ H : forall ctx Hwf', Named.interpf ?e = Some _, Hwf : prop_of_option (Named.wff _ ?e) |- _ ]
- => specialize (H _ Hwf)
- | [ H : forall ctx Hwf, Named.interpf ?e = Some _ |- Named.interpf ?e = Some _ ]
- => rewrite H by auto
- end ].
- Qed.
-
- Lemma interp_interp_to_phoas
- (ctx : Context)
- {t} (e : @Named.expr base_type_code op Name t)
- (Hwf : Named.wf ctx e)
- v
- : Named.interp (interp_op:=interp_op) (ctx:=ctx) e v
- = Some (Syntax.interp interp_op (interp_to_phoas failb ctx e) v).
- Proof using Type.
- unfold interp, interp_to_phoas, Named.interp; apply interpf_interpf_to_phoas; auto.
- Qed.
- End with_context.
-
- Section all.
- Context {Context : forall var, @Context base_type_code Name var}
- {ContextOk : forall var, ContextOk (Context var)}
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t)).
-
- Lemma Interp_InterpToPHOAS_gen
- {ctx : forall var, Context var}
- {t} (e : @Named.expr base_type_code op Name t)
- (Hwf : forall var, Named.wf (ctx var) e)
- v
- : Named.interp (interp_op:=interp_op) (ctx:=ctx _) e v
- = Some (Interp interp_op (InterpToPHOAS_gen failb ctx e) v).
- Proof using Type. apply interp_interp_to_phoas; auto. Qed.
-
- Lemma Interp_InterpToPHOAS
- {t} (e : @Named.expr base_type_code op Name t)
- (Hwf : Named.Wf Context e)
- v
- : Named.Interp (Context:=Context _) (interp_op:=interp_op) e v
- = Some (Interp interp_op (InterpToPHOAS (Context:=Context) failb e) v).
- Proof using Type. apply interp_interp_to_phoas; auto. Qed.
- End all.
-End language.
diff --git a/src/Compilers/Named/InterpretToPHOASWf.v b/src/Compilers/Named/InterpretToPHOASWf.v
deleted file mode 100644
index d2b3f18d3..000000000
--- a/src/Compilers/Named/InterpretToPHOASWf.v
+++ /dev/null
@@ -1,139 +0,0 @@
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.SmartMap.
-Require Import Crypto.Compilers.Named.InterpretToPHOAS.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {base_type_code_dec : DecidableRel (@eq base_type_code)}
- {Name_dec : DecidableRel (@eq Name)}.
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}
- {Context1 : Context Name var1}
- {Context2 : Context Name var2}
- {Context1Ok : ContextOk Context1}
- {Context2Ok : ContextOk Context2}
- (failb1 : forall t, @Syntax.exprf base_type_code op var1 (Tbase t))
- (failb2 : forall t, @Syntax.exprf base_type_code op var2 (Tbase t)).
-
- Local Ltac t_step :=
- first [ progress intros
- | progress unfold dec in *
- | reflexivity
- | progress subst
- | progress inversion_option
- | erewrite lookupb_extend by assumption
- | rewrite <- !find_Name_and_val_None_iff
- | progress break_innermost_match_step
- | progress break_match_hyps
- | solve [ eauto using find_Name_and_val_flatten_binding_list ]
- | congruence
- | tauto
- | match goal with
- | [ H : lookupb (extend _ _ _) _ = _ |- _ ]
- => erewrite (lookupb_extend _ _ _) in H by assumption
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => setoid_rewrite List.in_app_iff in H
- | [ |- context[List.In _ (_ ++ _)] ]
- => rewrite List.in_app_iff
- | [ |- context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default))
- | [ H : context[find_Name_and_val ?tdec ?ndec ?a ?b ?c ?d ?default] |- _ ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite (find_Name_and_val_split tdec ndec (default:=default)) in H
- | [ H : forall n t, lookupb _ n = None <-> lookupb _ n = None |- context[lookupb _ _ = None] ]
- => rewrite H
- | [ H : forall n t, lookupb _ n = None |- context[lookupb _ _ = None] ]
- => rewrite H
- end ].
- Local Ltac t := repeat t_step.
-
- Lemma wff_interpf_to_phoas
- (ctx1 : Context1) (ctx2 : Context2)
- {t} (e : @Named.exprf base_type_code op Name t)
- (Hwf1 : prop_of_option (Named.wff ctx1 e))
- (Hwf2 : prop_of_option (Named.wff ctx2 e))
- G
- (HG : forall n t v1 v2,
- lookupb t ctx1 n = Some v1
- -> lookupb t ctx2 n = Some v2
- -> List.In (existT _ t (v1, v2)%core) G)
- (Hctx1_ctx2 : forall n t,
- lookupb t ctx1 n = None <-> lookupb t ctx2 n = None)
- : wff G (interpf_to_phoas failb1 ctx1 e) (interpf_to_phoas failb2 ctx2 e).
- Proof using Context1Ok Context2Ok Name_dec base_type_code_dec.
- revert dependent G; revert dependent ctx1; revert dependent ctx2; induction e;
- repeat first [ progress intros
- | progress destruct_head' and
- | progress break_innermost_match_step
- | progress simpl in *
- | progress autorewrite with push_prop_of_option in *
- | solve [ eauto | tauto ]
- | match goal with
- | [ |- wff _ _ _ ] => constructor
- end ].
- match goal with H : _ |- _ => eapply H end; t.
- Qed.
-
- Lemma wf_interp_to_phoas_gen
- (ctx1 : Context1) (ctx2 : Context2)
- {t} (e : @Named.expr base_type_code op Name t)
- (Hwf1 : Named.wf ctx1 e)
- (Hwf2 : Named.wf ctx2 e)
- (Hctx1 : forall n t, lookupb t ctx1 n = None)
- (Hctx2 : forall n t, lookupb t ctx2 n = None)
- : wf (interp_to_phoas failb1 ctx1 e) (interp_to_phoas failb2 ctx2 e).
- Proof using Context1Ok Context2Ok Name_dec base_type_code_dec.
- constructor; intros.
- apply wff_interpf_to_phoas; t.
- Qed.
-
- Lemma wf_interp_to_phoas
- {t} (e : @Named.expr base_type_code op Name t)
- (Hwf1 : Named.wf (Context:=Context1) empty e)
- (Hwf2 : Named.wf (Context:=Context2) empty e)
- : wf (interp_to_phoas (Context:=Context1) failb1 empty e) (interp_to_phoas (Context:=Context2) failb2 empty e).
- Proof using Context1Ok Context2Ok Name_dec base_type_code_dec.
- apply wf_interp_to_phoas_gen; auto using lookupb_empty.
- Qed.
- End with_var.
-
- Section all.
- Context {Context : forall var, @Context base_type_code Name var}
- {ContextOk : forall var, ContextOk (Context var)}
- (failb : forall var t, @Syntax.exprf base_type_code op var (Tbase t)).
-
- Lemma Wf_InterpToPHOAS_gen
- {ctx : forall var, Context var}
- {t} (e : @Named.expr base_type_code op Name t)
- (Hctx : forall var n t, lookupb t (ctx var) n = None)
- (Hwf : forall var, Named.wf (ctx var) e)
- : Wf (InterpToPHOAS_gen failb ctx e).
- Proof using ContextOk Name_dec base_type_code_dec.
- intros ??; apply wf_interp_to_phoas_gen; auto.
- Qed.
-
- Lemma Wf_InterpToPHOAS
- {t} (e : @Named.expr base_type_code op Name t)
- (Hwf : Named.Wf Context e)
- : Wf (InterpToPHOAS (Context:=Context) failb e).
- Proof using ContextOk Name_dec base_type_code_dec.
- intros ??; apply wf_interp_to_phoas; auto.
- Qed.
- End all.
-End language.
-
-Hint Resolve Wf_InterpToPHOAS : wf.
diff --git a/src/Compilers/Named/MapCast.v b/src/Compilers/Named/MapCast.v
deleted file mode 100644
index fddee84fa..000000000
--- a/src/Compilers/Named/MapCast.v
+++ /dev/null
@@ -1,72 +0,0 @@
-Require Import Coq.Bool.Sumbool.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-
-Local Open Scope nexpr_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {interp_base_type_bounds : base_type_code -> 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_code).
- 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)))
- {BoundsContext : Context Name interp_base_type_bounds}.
-
- Fixpoint mapf_cast
- (ctx : BoundsContext)
- {t} (e : exprf base_type_code op Name t)
- {struct e}
- : option { bounds : interp_flat_type interp_base_type_bounds t
- & exprf base_type_code op Name (pick_type bounds) }
- := match e in exprf _ _ _ t return option { bounds : interp_flat_type interp_base_type_bounds t
- & exprf base_type_code op Name (pick_type bounds) } with
- | TT => Some (existT _ tt TT)
- | Pair tx ex ty ey
- => match @mapf_cast ctx _ ex, @mapf_cast ctx _ ey with
- | Some (existT x_bs xv), Some (existT y_bs yv)
- => Some (existT _ (x_bs, y_bs)%core (Pair xv yv))
- | None, _ | _, None => None
- end
- | Var t x
- => option_map
- (fun bounds => existT _ bounds (Var x))
- (lookupb (t:=t) ctx x)
- | LetIn tx n ex tC eC
- => match @mapf_cast ctx _ ex with
- | Some (existT x_bounds ex')
- => option_map
- (fun eC' => let 'existT Cx_bounds C_expr := eC' in
- existT _ Cx_bounds (LetIn (pick_type x_bounds)
- (SmartFlatTypeMapInterp2 (t:=tx) (fun _ _ (n : Name) => n) x_bounds n) ex' C_expr))
- (@mapf_cast (extend (t:=tx) ctx n x_bounds) _ eC)
- | None => None
- end
- | Op t tR opc args
- => option_map
- (fun args'
- => let 'existT args_bounds argsv := args' in
- existT _
- (interp_op_bounds _ _ _ args_bounds)
- (Op (cast_op t tR opc args_bounds) argsv))
- (@mapf_cast ctx _ args)
- end.
-
- Definition map_cast
- (ctx : BoundsContext)
- {t} (e : expr base_type_code op Name t)
- (input_bounds : interp_flat_type interp_base_type_bounds (domain t))
- : option { output_bounds : interp_flat_type interp_base_type_bounds (codomain t)
- & expr base_type_code op Name (Arrow (pick_type input_bounds) (pick_type output_bounds)) }
- := option_map
- (fun v => existT
- _
- (projT1 v)
- (Abs (SmartFlatTypeMapInterp2 (fun _ _ (n' : Name) => n') input_bounds (Abs_name e))
- (projT2 v)))
- (mapf_cast (extend ctx (Abs_name e) input_bounds) (invert_Abs e)).
-End language.
diff --git a/src/Compilers/Named/MapCastInterp.v b/src/Compilers/Named/MapCastInterp.v
deleted file mode 100644
index c38eadd9c..000000000
--- a/src/Compilers/Named/MapCastInterp.v
+++ /dev/null
@@ -1,290 +0,0 @@
-Require Import Coq.Bool.Sumbool.
-Require Import Coq.Logic.Eqdep_dec.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.SmartMap.
-Require Import Crypto.Compilers.Named.InterpSideConditions.
-Require Import Crypto.Compilers.Named.InterpSideConditionsInterp.
-Require Import Crypto.Compilers.Named.MapCast.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-
-Local Open Scope nexpr_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {interp_base_type_bounds : base_type_code -> 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_code).
- Local Notation pick_type t v := (SmartFlatTypeMap pick_typeb (t:=t) 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)))
- {BoundsContext : Context Name interp_base_type_bounds}
- (BoundsContextOk : ContextOk BoundsContext)
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst,
- op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst)
- (interped_op_side_conditions : forall s d, op s d -> interp_flat_type interp_base_type s -> pointed_Prop)
- (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 t b) -> interp_flat_type interp_base_type t
- := fun t b => SmartFlatTypeMapUnInterp cast_backb.
- Context {Context : Context Name interp_base_type}
- (ContextOk : ContextOk 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)
- (Hside : to_prop (interped_op_side_conditions _ _ opc 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 t bs))
- (H : inbounds t bs (cast_back t bs v))
- (Hside : to_prop (interped_op_side_conditions _ _ opc (cast_back t bs v))),
- interp_op t tR opc (cast_back t bs v)
- =
- cast_back _ _ (interp_op _ _ (cast_op _ _ opc bs) v))
- (base_type_dec : DecidableRel (@eq base_type_code))
- (Name_dec : DecidableRel (@eq Name)).
-
- Local Notation mapf_cast := (@mapf_cast _ op Name _ interp_op_bounds pick_typeb cast_op BoundsContext).
- Local Notation map_cast := (@map_cast _ op Name _ interp_op_bounds pick_typeb cast_op BoundsContext).
-
- Local Ltac handle_options_step :=
- match goal with
- | _ => progress inversion_option
- | [ H : ?x = Some _ |- context[?x] ] => rewrite H
- | [ H : ?x = None |- context[?x] ] => rewrite H
- | [ H : ?x = Some _, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : ?x = None, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : Some _ <> None \/ _ |- _ ] => clear H
- | [ H : Some ?x <> Some ?y |- _ ] => assert (x <> y) by congruence; clear H
- | [ H : None <> Some _ |- _ ] => clear H
- | [ H : Some _ <> None |- _ ] => clear H
- | [ H : ?x <> ?x \/ _ |- _ ] => destruct H; [ exfalso; apply H; reflexivity | ]
- | [ H : _ \/ None = Some _ |- _ ] => destruct H; [ | exfalso; clear -H; congruence ]
- | [ H : _ \/ Some _ = None |- _ ] => destruct H; [ | exfalso; clear -H; congruence ]
- | [ H : ?x = Some ?y, H' : ?x = Some ?y' |- _ ]
- => assert (y = y') by congruence; (subst y' || subst y)
- | _ => progress simpl @option_map
- end.
-
- Local Ltac handle_lookupb_step :=
- let do_eq_dec dec t t' :=
- first [ constr_eq t t'; fail 1
- | lazymatch goal with
- | [ H : t = t' |- _ ] => fail 1
- | [ H : t <> t' |- _ ] => fail 1
- | [ H : t = t' -> False |- _ ] => fail 1
- | _ => destruct (dec t t')
- end ] in
- let do_type_dec := do_eq_dec base_type_dec in
- match goal with
- | _ => progress unfold dec in *
- | _ => handle_options_step
- (* preprocess *)
- | [ H : context[lookupb (extend _ _ _) _] |- _ ]
- => first [ rewrite (lookupb_extend base_type_dec Name_dec) in H by assumption
- | setoid_rewrite (lookupb_extend base_type_dec Name_dec) in H; [ | assumption.. ] ]
- | [ |- context[lookupb (extend _ _ _) _] ]
- => first [ rewrite (fun C => lookupb_extend C base_type_dec Name_dec) by assumption
- | setoid_rewrite (lookupb_extend base_type_dec Name_dec); [ | assumption.. ] ]
- | _ => progress subst
- (* handle multiple hypotheses *)
- | [ H : find_Name _ ?n ?N = Some ?t', H'' : context[find_Name_and_val _ _ ?t ?n ?N ?x ?default] |- _ ]
- => do_type_dec t t'
- (* clear the default value *)
- | [ H : context[find_Name_and_val ?tdec ?ndec ?t ?n (T:=?T) ?N ?V ?default] |- _ ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite find_Name_and_val_split in H
- (* generic handlers *)
- | [ H : find_Name _ ?n ?N = Some ?t', H' : ?t <> ?t', H'' : context[find_Name_and_val _ _ ?t ?n ?N ?x ?default] |- _ ]
- => erewrite find_Name_and_val_wrong_type in H'' by eassumption
- | [ H : context[find_Name _ _ (SmartFlatTypeMapInterp2 _ _ _)] |- _ ]
- => rewrite find_Name_SmartFlatTypeMapInterp2 with (base_type_code_dec:=base_type_dec) in H
- | [ H : find_Name_and_val _ _ _ _ _ _ _ = None |- _ ]
- => apply find_Name_and_val_None_iff in H
- (* destructers *)
- | [ |- context[find_Name_and_val ?tdec ?ndec ?t ?n ?N ?V ?default] ]
- => destruct (find_Name_and_val tdec ndec t n N V default) eqn:?
- | [ H : context[match find_Name_and_val ?tdec ?ndec ?t ?n ?N ?V ?default with _ => _ end] |- _ ]
- => destruct (find_Name_and_val tdec ndec t n N V default) eqn:?
- | [ H : context[match find_Name ?ndec ?n ?N with _ => _ end] |- _ ]
- => destruct (find_Name ndec n N) eqn:?
- | [ H : context[match base_type_dec ?x ?y with _ => _ end] |- _ ]
- => destruct (base_type_dec x y)
- | [ H : context[match Name_dec ?x ?y with _ => _ end] |- _ ]
- => destruct (Name_dec x y)
- end.
-
- Local Ltac handle_exists_in_goal :=
- lazymatch goal with
- | [ |- exists v, Some ?k = Some v /\ @?B v ]
- => exists k; split; [ reflexivity | ]
- | [ |- (exists v, None = Some v /\ @?B v) ]
- => exfalso
- | [ |- ?A /\ (exists v, Some ?k = Some v /\ @?B v) ]
- => cut (A /\ B k); [ clear; solve [ intuition eauto ] | cbv beta ]
- | [ |- ?A /\ (exists v, None = Some v /\ @?B v) ]
- => exfalso
- end.
- Local Ltac handle_bounds_side_conditions_step :=
- match goal with
- | [ H : interpf ?e = Some ?v, H' : interpf_side_conditions_gen _ _ _ ?e = Some (_, ?v')%core |- _ ]
- => first [ constr_eq v v'; fail 1
- | assert (Some v = Some v')
- by (erewrite <- H, snd_interpf_side_conditions_gen_eq, H'; reflexivity);
- inversion_option; (subst v || subst v') ]
- end.
- Local Ltac fin_inbounds_cast_back_t_step :=
- match goal with
- | [ |- inboundsb _ _ _ /\ _ ]
- => split; [ eapply interp_flat_type_rel_pointwise__find_Name_and_val; eassumption | ]
- | [ |- cast_backb _ _ _ = _ ]
- => eapply find_Name_and_val__SmartFlatTypeMapInterp2__SmartFlatTypeMapUnInterp__Some_Some; [ | eassumption.. ]
- end.
- Local Ltac specializer_t_step :=
- match goal with
- | [ H : ?T, H' : ?T |- _ ] => clear H
- | [ H : forall x, Some _ = Some x -> _ |- _ ] => specialize (H _ eq_refl)
- | [ H : ?x = Some _, IH : forall a b, ?x = Some _ -> _ |- _ ]
- => specialize (IH _ _ H)
- | [ H : ?x = Some _, IH : forall a, ?x = Some _ -> _ |- _ ]
- => specialize (IH _ H)
- | [ H : forall t n v, lookupb ?ctx n = _ -> _, H' : lookupb ?ctx ?n' = _ |- _ ]
- => specialize (H _ _ _ H')
- | _ => progress specialize_by auto
- end.
-
- Local Ltac break_t_step :=
- first [ progress destruct_head'_ex
- | progress destruct_head'_and ].
-
- Local Ltac t_step :=
- first [ progress intros
- | break_t_step
- | handle_lookupb_step
- | handle_exists_in_goal
- | solve [ auto ]
- | specializer_t_step
- | fin_inbounds_cast_back_t_step
- | handle_options_step
- | handle_bounds_side_conditions_step ].
- Local Ltac t := repeat t_step.
-
- Local Ltac do_specialize_IHe :=
- repeat match goal with
- | [ IH : context[interpf ?e], H' : interpf (ctx:=?ctx) ?e = _ |- _ ]
- => let check_tac _ := (rewrite H' in IH) in
- first [ specialize (IH ctx); check_tac ()
- | specialize (fun a => IH a ctx); check_tac ()
- | specialize (fun a b => IH a b ctx); check_tac () ]
- | [ IH : context[mapf_cast _ ?e], H' : mapf_cast ?ctx ?e = _ |- _ ]
- => let check_tac _ := (rewrite H' in IH) in
- first [ specialize (IH ctx); check_tac ()
- | specialize (fun a => IH a ctx); check_tac ()
- | specialize (fun a b => IH a b ctx); check_tac () ]
- | [ H : forall x y z, Some _ = Some _ -> _ |- _ ]
- => first [ specialize (H _ _ _ eq_refl)
- | specialize (fun x => H x _ _ eq_refl) ]
- | [ H : forall x y, Some _ = Some _ -> _ |- _ ]
- => first [ specialize (H _ _ eq_refl)
- | specialize (fun x => H x _ eq_refl) ]
- | _ => progress specialize_by_assumption
- end.
-
- Lemma mapf_cast_correct
- {t} (e:exprf base_type_code op Name t)
- : forall
- (oldValues:Context)
- (newValues:Context)
- (varBounds:BoundsContext)
- {b} e' (He':mapf_cast varBounds e = Some (existT _ b e'))
- (Hctx:forall {t} n v,
- lookupb (t:=t) oldValues n = Some v
- -> exists b, lookupb (t:=t) varBounds n = Some b
- /\ @inboundsb _ b v
- /\ exists v', lookupb (t:=pick_typeb t b) newValues n = Some v'
- /\ cast_backb t b v' = v)
- r (Hr:interpf (interp_op:=interp_op) (ctx:=oldValues) e = Some r)
- r' (Hr':interpf (interp_op:=interp_op) (ctx:=newValues) e' = Some r')
- (Hside : prop_of_option (interpf_side_conditions interp_op interped_op_side_conditions oldValues e))
- , interpf (interp_op:=interp_op_bounds) (ctx:=varBounds) e = Some b
- /\ @inbounds _ b r /\ cast_back _ _ r' = r.
- Proof using Type*.
- induction e; simpl interpf; simpl mapf_cast; unfold option_map, cast_back in *; intros;
- unfold interpf_side_conditions in *; simpl in Hside;
- repeat (repeat handle_options_step; break_match_hyps; inversion_option; inversion_sigma; autorewrite with push_to_prop in *; simpl in *; unfold option_map in *; subst; try tauto).
- { destruct (Hctx _ _ _ Hr) as [b' [Hb'[Hb'v[v'[Hv' Hv'v]]]]]; clear Hctx Hr; subst.
- repeat match goal with
- [H: ?e = Some ?x, G:?e = Some ?x' |- _] =>
- pose proof (eq_trans (eq_sym G) H); clear G; inversion_option; subst
- end.
- auto. }
- { do_specialize_IHe.
- repeat (handle_options_step || destruct_head_and || specialize_by_assumption).
- subst; intuition eauto; try (symmetry; rewrite_hyp ?*; eauto);
- repeat handle_bounds_side_conditions_step; auto. }
- { cbv [LetIn.Let_In] in *.
- do_specialize_IHe.
- destruct_head'_and.
- destruct IHe1 as [IHe1_eq IHe1]; rewrite_hyp *; try assumption.
- { apply IHe2; clear IHe2; try reflexivity; [ | t ].
- intros ??? Hlookup.
- let b := fresh "b" in
- let H' := fresh "H'" in
- match goal with |- exists b0, ?v = Some b0 /\ _ => destruct v as [b|] eqn:H' end;
- [ exists b; split; [ reflexivity | ] | exfalso ];
- revert Hlookup H'; t. } }
- { do_specialize_IHe.
- t. }
- Qed.
-
- Lemma map_cast_correct
- {t} (e:expr base_type_code op Name t)
- (input_bounds : interp_flat_type interp_base_type_bounds (domain t))
- : forall
- (oldValues:Context)
- (newValues:Context)
- (varBounds:BoundsContext)
- {b} e' (He':map_cast varBounds e input_bounds = Some (existT _ b e'))
- (Hctx:forall {t} n v,
- lookupb (t:=t) oldValues n = Some v
- -> exists b, lookupb (t:=t) varBounds n = Some b
- /\ @inboundsb _ b v
- /\ exists v', lookupb (t:=pick_typeb t b) newValues n = Some v'
- /\ cast_backb t b v' = v)
- v v' (Hv : @inbounds _ input_bounds v /\ cast_back _ _ v' = v)
- r (Hr:interp (interp_op:=interp_op) (ctx:=oldValues) e v = Some r)
- r' (Hr':interp (interp_op:=interp_op) (ctx:=newValues) e' v' = Some r')
- (Hside : prop_of_option (interp_side_conditions interp_op interped_op_side_conditions oldValues e v))
- , interp (interp_op:=interp_op_bounds) (ctx:=varBounds) e input_bounds = Some b
- /\ @inbounds _ b r /\ cast_back _ _ r' = r.
- Proof using Type*.
- unfold map_cast, option_map, interp; simpl; intros.
- repeat first [ progress subst
- | progress inversion_option
- | progress inversion_sigma
- | progress break_match_hyps
- | progress destruct_head' sigT
- | progress simpl in * ].
- eapply mapf_cast_correct; try eassumption.
- t.
- Qed.
-End language.
diff --git a/src/Compilers/Named/MapCastWf.v b/src/Compilers/Named/MapCastWf.v
deleted file mode 100644
index eb141cad1..000000000
--- a/src/Compilers/Named/MapCastWf.v
+++ /dev/null
@@ -1,285 +0,0 @@
-Require Import Coq.Bool.Sumbool.
-Require Import Coq.Logic.Eqdep_dec.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextProperties.SmartMap.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Compilers.Named.MapCast.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Local Open Scope nexpr_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}
- {interp_base_type_bounds : base_type_code -> 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_code).
- Local Notation pick_type t v := (SmartFlatTypeMap pick_typeb (t:=t) 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)))
- {BoundsContext : Context Name interp_base_type_bounds}
- (BoundsContextOk : ContextOk BoundsContext)
- {interp_base_type : base_type_code -> Type}
- (interp_op : forall src dst,
- op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst)
- {FullContext : Context Name (fun t => { b : interp_base_type_bounds t & interp_base_type (pick_typeb t b) }%type)}
- (FullContextOk : ContextOk FullContext)
- {Context : Context Name interp_base_type}
- (ContextOk : ContextOk Context)
- (base_type_dec : DecidableRel (@eq base_type_code))
- (Name_dec : DecidableRel (@eq Name)).
-
- Local Notation mapf_cast := (@mapf_cast _ op Name _ interp_op_bounds pick_typeb cast_op BoundsContext).
- Local Notation map_cast := (@map_cast _ op Name _ interp_op_bounds pick_typeb cast_op BoundsContext).
-
- Local Ltac handle_options_step :=
- match goal with
- | _ => progress inversion_option
- | [ H : ?x = Some _ |- context[?x] ] => rewrite H
- | [ H : ?x = None |- context[?x] ] => rewrite H
- | [ H : ?x = Some _, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : ?x = None, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : Some _ <> None \/ _ |- _ ] => clear H
- | [ H : Some ?x <> Some ?y |- _ ] => assert (x <> y) by congruence; clear H
- | [ H : None <> Some _ |- _ ] => clear H
- | [ H : Some _ <> None |- _ ] => clear H
- | [ H : ?x <> ?x \/ _ |- _ ] => destruct H; [ exfalso; apply H; reflexivity | ]
- | [ H : _ \/ None = Some _ |- _ ] => destruct H; [ | exfalso; clear -H; congruence ]
- | [ H : _ \/ Some _ = None |- _ ] => destruct H; [ | exfalso; clear -H; congruence ]
- | [ H : ?x = Some ?y, H' : ?x = Some ?y' |- _ ]
- => assert (y = y') by congruence; (subst y' || subst y)
- | _ => progress simpl @option_map
- | _ => progress unfold option_map in *
- end.
-
- Local Ltac handle_lookupb_step_extra := fail.
- Local Ltac handle_lookupb_step :=
- let do_eq_dec dec t t' :=
- first [ constr_eq t t'; fail 1
- | lazymatch goal with
- | [ H : t = t' |- _ ] => fail 1
- | [ H : t <> t' |- _ ] => fail 1
- | [ H : t = t' -> False |- _ ] => fail 1
- | _ => destruct (dec t t')
- end ] in
- let do_type_dec := do_eq_dec base_type_dec in
- match goal with
- | _ => progress unfold dec in *
- | _ => handle_options_step
- (* preprocess *)
- | [ H : context[lookupb (extend _ _ _) _] |- _ ]
- => first [ rewrite (lookupb_extend base_type_dec Name_dec) in H by assumption
- | setoid_rewrite (lookupb_extend base_type_dec Name_dec) in H; [ | assumption.. ] ]
- | [ |- context[lookupb (extend _ _ _) _] ]
- => first [ rewrite (lookupb_extend base_type_dec Name_dec) by assumption
- | setoid_rewrite (lookupb_extend base_type_dec Name_dec); [ | assumption.. ] ]
- | _ => progress subst
- (* handle multiple hypotheses *)
- | [ H : find_Name _ ?n ?N = Some ?t', H'' : context[find_Name_and_val _ _ ?t ?n ?N ?x ?default] |- _ ]
- => do_type_dec t t'
- (* clear the default value *)
- | [ H : context[find_Name_and_val ?tdec ?ndec ?t ?n (T:=?T) ?N ?V ?default] |- _ ]
- => lazymatch default with None => fail | _ => idtac end;
- rewrite find_Name_and_val_split in H
- (* generic handlers *)
- | [ H : find_Name _ ?n ?N = Some ?t', H' : ?t <> ?t', H'' : context[find_Name_and_val _ _ ?t ?n ?N ?x ?default] |- _ ]
- => erewrite find_Name_and_val_wrong_type in H'' by eassumption
- | [ H : context[find_Name _ _ (SmartFlatTypeMapInterp2 _ _ _)] |- _ ]
- => rewrite find_Name_SmartFlatTypeMapInterp2 with (base_type_code_dec:=base_type_dec) in H
- | [ H : find_Name_and_val _ _ _ _ _ _ _ = None |- _ ]
- => apply find_Name_and_val_None_iff in H
- | _ => progress handle_lookupb_step_extra
- (* destructers *)
- | [ |- context[find_Name_and_val ?tdec ?ndec ?t ?n ?N ?V ?default] ]
- => destruct (find_Name_and_val tdec ndec t n N V default) eqn:?
- | [ H : context[match find_Name_and_val ?tdec ?ndec ?t ?n ?N ?V ?default with _ => _ end] |- _ ]
- => destruct (find_Name_and_val tdec ndec t n N V default) eqn:?
- | [ H : context[match find_Name ?ndec ?n ?N with _ => _ end] |- _ ]
- => destruct (find_Name ndec n N) eqn:?
- | [ H : context[match base_type_dec ?x ?y with _ => _ end] |- _ ]
- => destruct (base_type_dec x y)
- | [ H : context[match Name_dec ?x ?y with _ => _ end] |- _ ]
- => destruct (Name_dec x y)
- end.
-
- Local Ltac handle_exists_in_goal :=
- lazymatch goal with
- | [ |- exists v, Some ?k = Some v /\ @?B v ]
- => exists k; split; [ reflexivity | ]
- | [ |- exists v, Some ?k = Some v ]
- => exists k; reflexivity
- | [ |- (exists v, None = Some v /\ @?B v) ]
- => exfalso
- | [ |- ?A /\ (exists v, Some ?k = Some v /\ @?B v) ]
- => cut (A /\ B k); [ clear; solve [ intuition eauto ] | cbv beta ]
- | [ |- ?A /\ (exists v, None = Some v /\ @?B v) ]
- => exfalso
- end.
- Local Ltac specializer_t_step :=
- match goal with
- | [ H : ?T, H' : ?T |- _ ] => clear H
- | [ H : forall x, Some _ = Some x -> _ |- _ ] => specialize (H _ eq_refl)
- | [ H : ?x = Some _, IH : forall a b c, ?x = Some _ -> _ |- _ ]
- => specialize (IH _ _ _ H)
- | [ H : ?x = Some _, IH : forall a b, ?x = Some _ -> _ |- _ ]
- => specialize (IH _ _ H)
- | [ H : ?x = Some _, IH : forall a, ?x = Some _ -> _ |- _ ]
- => specialize (IH _ H)
- | [ H : forall t n x y z, lookupb ?ctx n = _ -> _, H' : lookupb ?ctx ?n' = _ |- _ ]
- => specialize (H _ _ _ _ _ H')
- | [ H : forall t n x y, lookupb ?ctx n = _ -> _, H' : lookupb ?ctx ?n' = _ |- _ ]
- => specialize (H _ _ _ _ H')
- | [ H : forall t n v, lookupb ?ctx n = _ -> _, H' : lookupb ?ctx ?n' = _ |- _ ]
- => specialize (H _ _ _ H')
- | _ => progress specialize_by auto
- end.
-
- Local Ltac break_t_step :=
- first [ progress subst
- | progress destruct_head'_ex
- | progress destruct_head'_and
- | progress inversion_option
- | progress inversion_prod
- | progress inversion_sigma
- | progress autorewrite with push_prop_of_option in *
- | progress break_match_hyps ].
-
- Local Ltac do_specialize_IHe_step :=
- match goal with
- | [ IH : context[mapf_cast _ ?e], H' : mapf_cast ?ctx ?e = _ |- _ ]
- => let check_tac _ := (rewrite H' in IH) in
- first [ specialize (IH ctx); check_tac ()
- | specialize (fun a => IH a ctx); check_tac ()
- | specialize (fun a b => IH a b ctx); check_tac () ]
- | [ H : forall x y z w, Some _ = Some _ -> _ |- _ ]
- => first [ specialize (H _ _ _ _ eq_refl)
- | specialize (fun x y => H x y _ _ eq_refl) ]
- | [ H : forall x y z, Some _ = Some _ -> _ |- _ ]
- => first [ specialize (H _ _ _ eq_refl)
- | specialize (fun x => H x _ _ eq_refl) ]
- | [ H : forall x y, Some _ = Some _ -> _ |- _ ]
- => first [ specialize (H _ _ eq_refl)
- | specialize (fun x => H x _ eq_refl) ]
- | _ => progress specialize_by_assumption
- | [ H : forall a b, prop_of_option (Named.wff a ?e) -> _, H' : prop_of_option (Named.wff _ ?e) |- _ ]
- => specialize (fun b => H _ b H')
- | [ H : forall b v, _ -> prop_of_option (Named.wff b ?e) |- prop_of_option (Named.wff ?ctx ?e) ]
- => specialize (H ctx)
- | [ H : forall b v, _ -> _ -> prop_of_option (Named.wff b ?e) |- prop_of_option (Named.wff ?ctx ?e) ]
- => specialize (H ctx)
- | [ H : forall a b, _ -> _ -> _ -> prop_of_option (Named.wff b ?e) |- prop_of_option (Named.wff ?ctx ?e) ]
- => specialize (fun a => H a ctx)
- | [ H : forall a b, prop_of_option (Named.wff a ?e) -> _, H' : forall v, prop_of_option (Named.wff _ ?e) |- _ ]
- => specialize (fun b v => H _ b (H' v))
- end.
- Ltac do_specialize_IHe := repeat do_specialize_IHe_step.
-
- Definition make_fContext_value {t} {b : interp_flat_type interp_base_type_bounds t}
- (v : interp_flat_type interp_base_type (pick_type t b))
- : interp_flat_type
- (fun t => { b : interp_base_type_bounds t & interp_base_type (pick_typeb t b)})
- t
- := SmartFlatTypeMapUnInterp2
- (fun t b (v : interp_flat_type _ (Tbase _))
- => existT (fun b => interp_base_type (pick_typeb t b)) b v)
- v.
-
- Local Ltac t_step :=
- first [ progress intros
- | progress simpl in *
- | break_t_step
- | handle_lookupb_step
- | handle_exists_in_goal
- | apply conj
- | solve [ auto | exfalso; auto ]
- | specializer_t_step
- | progress do_specialize_IHe
- | match goal with
- | [ IH : forall v, _ -> ?T, v' : interp_flat_type _ _ |- ?T ]
- => apply (IH (make_fContext_value v')); clear IH
- end ].
- Local Ltac t := repeat t_step.
-
- Lemma find_Name_and_val_make_fContext_value_Some {T}
- {N : interp_flat_type (fun _ : base_type_code => Name) T}
- {B : interp_flat_type interp_base_type_bounds T}
- {V : interp_flat_type interp_base_type (pick_type T B)}
- {n : Name}
- {t : base_type_code}
- {v : { b : interp_base_type_bounds t & interp_base_type (pick_typeb t b)}}
- {b}
- (Hn : find_Name Name_dec n N = Some t)
- (Hf : find_Name_and_val base_type_dec Name_dec t n N (make_fContext_value V) None = Some v)
- (Hb : find_Name_and_val base_type_dec Name_dec t n N B None = Some b)
- (N' := SmartFlatTypeMapInterp2 (var'':=fun _ => Name) (f:=pick_typeb) (fun _ _ n => n) _ N)
- : b = projT1 v /\ find_Name_and_val base_type_dec Name_dec (pick_typeb t (projT1 v)) n N' V None = Some (projT2 v).
- Proof using Type.
- eapply (find_Name_and_val_SmartFlatTypeMapUnInterp2_Some_Some base_type_dec Name_dec (h:=@projT1 _ _) (i:=@projT2 _ _) (f:=pick_typeb) (g:=fun _ => existT _));
- auto.
- Qed.
-
- Local Ltac handle_lookupb_step_extra ::=
- lazymatch goal with
- | [ H : find_Name _ ?n ?N = Some ?t,
- H' : find_Name_and_val _ _ ?t ?n ?N (@make_fContext_value ?T ?B ?v) None = Some ?v',
- H'' : find_Name_and_val _ _ ?t ?n ?N ?B None = Some _
- |- _ ]
- => pose proof (find_Name_and_val_make_fContext_value_Some H H' H''); clear H'
- end.
-
- Lemma wff_mapf_cast
- {t} (e:exprf base_type_code op Name t)
- : forall
- (fValues:FullContext)
- (newValues:Context)
- (varBounds:BoundsContext)
- {b} e' (He':mapf_cast varBounds e = Some (existT _ b e'))
- (Hwf : prop_of_option (Named.wff fValues e))
- (Hctx:forall {t} n v,
- lookupb (t:=t) fValues n = Some v
- -> lookupb (t:=t) varBounds n = Some (projT1 v)
- /\ lookupb (t:=pick_typeb t (projT1 v)) newValues n = Some (projT2 v)),
- prop_of_option (Named.wff newValues e').
- Proof using BoundsContextOk ContextOk FullContextOk Name_dec base_type_dec. induction e; t. Qed.
-
- Lemma wf_map_cast
- {t} (e:expr base_type_code op Name t)
- (input_bounds : interp_flat_type interp_base_type_bounds (domain t))
- : forall
- (fValues:FullContext)
- (newValues:Context)
- (varBounds:BoundsContext)
- {b} e' (He':map_cast varBounds e input_bounds = Some (existT _ b e'))
- (Hwf : Named.wf fValues e)
- (Hctx:forall {t} n v,
- lookupb (t:=t) fValues n = Some v
- -> lookupb (t:=t) varBounds n = Some (projT1 v)
- /\ lookupb (t:=pick_typeb t (projT1 v)) newValues n = Some (projT2 v)),
- Named.wf newValues e'.
- Proof using BoundsContextOk ContextOk FullContextOk Name_dec base_type_dec.
- unfold Named.wf, map_cast, option_map, interp; simpl; intros.
- repeat first [ progress subst
- | progress inversion_option
- | progress inversion_sigma
- | progress break_match_hyps
- | progress destruct_head' sigT
- | progress simpl in * ].
- match goal with v : _ |- _ => specialize (Hwf (make_fContext_value v)) end.
- eapply wff_mapf_cast; eauto; [].
- t.
- Qed.
-End language.
diff --git a/src/Compilers/Named/MapType.v b/src/Compilers/Named/MapType.v
deleted file mode 100644
index 15b40f7b7..000000000
--- a/src/Compilers/Named/MapType.v
+++ /dev/null
@@ -1,59 +0,0 @@
-Require Import Coq.Bool.Sumbool.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-
-Local Open Scope nexpr_scope.
-Section language.
- Context {base_type_code1 base_type_code2 : Type}
- {op1 : flat_type base_type_code1 -> flat_type base_type_code1 -> Type}
- {op2 : flat_type base_type_code2 -> flat_type base_type_code2 -> Type}
- {Name : Type}
- (f_base : base_type_code1 -> base_type_code2)
- (f_op : forall s d,
- op1 s d
- -> option (op2 (lift_flat_type f_base s) (lift_flat_type f_base d))).
-
- Fixpoint mapf_type
- {t} (e : exprf base_type_code1 op1 Name t)
- : option (exprf base_type_code2 op2 Name (lift_flat_type f_base t))
- := match e in exprf _ _ _ t return option (exprf _ _ _ (lift_flat_type f_base t)) with
- | TT => Some TT
- | Var t x => Some (Var x)
- | Op t1 tR opc args
- => let opc := f_op _ _ opc in
- let args := @mapf_type _ args in
- match opc, args with
- | Some opc, Some args => Some (Op opc args)
- | _, _ => None
- end
- | LetIn tx n ex tC eC
- => let n := transfer_interp_flat_type f_base (fun _ (n : Name) => n) n in
- let ex := @mapf_type _ ex in
- let eC := @mapf_type _ eC in
- match ex, eC with
- | Some ex, Some eC
- => Some (LetIn _ n ex eC)
- | _, _ => None
- end
- | Pair tx ex ty ey
- => let ex := @mapf_type _ ex in
- let ey := @mapf_type _ ey in
- match ex, ey with
- | Some ex, Some ey
- => Some (Pair ex ey)
- | _, _ => None
- end
- end.
-
- Definition map_type
- {t} (e : expr base_type_code1 op1 Name t)
- : option (expr base_type_code2 op2 Name (Arrow (lift_flat_type f_base (domain t)) (lift_flat_type f_base (codomain t))))
- := let f := invert_Abs e in
- let f := mapf_type f in
- let n := Abs_name e in
- let n := transfer_interp_flat_type f_base (fun _ (n : Name) => n) n in
- option_map
- (fun f => Abs n f)
- f.
-End language.
diff --git a/src/Compilers/Named/NameUtil.v b/src/Compilers/Named/NameUtil.v
deleted file mode 100644
index 8b099ffdf..000000000
--- a/src/Compilers/Named/NameUtil.v
+++ /dev/null
@@ -1,56 +0,0 @@
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Syntax.
-
-Local Open Scope core_scope.
-Local Notation eta x := (fst x, snd x).
-
-Section language.
- Context {base_type_code : Type}
- {Name : Type}.
-
- Section monad.
- Context (MName : Type) (force : MName -> option Name).
- Fixpoint split_mnames
- (t : flat_type base_type_code) (ls : list MName)
- : option (interp_flat_type (fun _ => Name) t) * list MName
- := match t return option (@interp_flat_type base_type_code (fun _ => Name) t) * _ with
- | Tbase _
- => match ls with
- | cons n ls'
- => match force n with
- | Some n => (Some n, ls')
- | None => (None, ls')
- end
- | nil => (None, nil)
- end
- | Unit => (Some tt, ls)
- | Prod A B
- => let '(a, ls) := eta (@split_mnames A ls) in
- let '(b, ls) := eta (@split_mnames B ls) in
- (match a, b with
- | Some a', Some b' => Some (a', b')
- | _, _ => None
- end,
- ls)
- end.
- Definition mname_list_unique (ls : list MName) : Prop
- := forall k n,
- List.In (Some n) (firstn k (List.map force ls))
- -> List.In (Some n) (skipn k (List.map force ls))
- -> False.
- End monad.
- Definition split_onames := @split_mnames (option Name) (fun x => x).
- Definition split_names := @split_mnames Name (@Some _).
-
- Definition oname_list_unique (ls : list (option Name)) : Prop
- := mname_list_unique (option Name) (fun x => x) ls.
- Definition name_list_unique (ls : list Name) : Prop
- := oname_list_unique (List.map (@Some Name) ls).
-End language.
-
-Global Arguments split_mnames {_ _ MName} force _ _, {_ _} MName force _ _.
-Global Arguments split_onames {_ _} _ _.
-Global Arguments split_names {_ _} _ _.
-Global Arguments mname_list_unique {_ MName} force ls, {_} MName force ls.
-Global Arguments oname_list_unique {_} ls.
-Global Arguments name_list_unique {_} ls.
diff --git a/src/Compilers/Named/NameUtilProperties.v b/src/Compilers/Named/NameUtilProperties.v
deleted file mode 100644
index 944d164f2..000000000
--- a/src/Compilers/Named/NameUtilProperties.v
+++ /dev/null
@@ -1,223 +0,0 @@
-Require Import Coq.omega.Omega.
-Require Import Coq.Arith.Arith.
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.CountLets.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.NatUtil.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-
-Local Open Scope core_scope.
-Section language.
- Context {base_type_code : Type}
- {Name : Type}.
-
- Section monad.
- Context (MName : Type) (force : MName -> option Name).
-
- Lemma split_mnames_firstn_skipn
- (t : flat_type base_type_code) (ls : list MName)
- : split_mnames force t ls
- = (fst (split_mnames force t (firstn (count_pairs t) ls)),
- skipn (count_pairs t) ls).
- Proof using Type.
- apply path_prod_uncurried; simpl.
- revert ls; induction t; split; split_prod;
- repeat first [ progress simpl in *
- | progress intros
- | rewrite <- skipn_skipn
- | reflexivity
- | progress break_innermost_match_step
- | apply (f_equal2 (@pair _ _))
- | rewrite_hyp <- !*
- | match goal with
- | [ H : forall ls, snd (split_mnames _ _ _) = _, H' : context[snd (split_mnames _ _ _)] |- _ ]
- => rewrite H in H'
- | [ H : _ |- _ ] => first [ rewrite <- firstn_skipn_add in H ]
- | [ H : forall ls', fst (split_mnames _ _ _) = _, H' : context[fst (split_mnames _ _ (skipn ?n ?ls))] |- _ ]
- => rewrite (H (skipn n ls)) in H'
- | [ H : forall ls', fst (split_mnames _ _ _) = _, H' : context[fst (split_mnames _ ?t (firstn (count_pairs ?t + ?n) ?ls))] |- _ ]
- => rewrite (H (firstn (count_pairs t + n) ls)), firstn_firstn in H' by omega
- | [ H : forall ls', fst (split_mnames _ _ _) = _, H' : context[fst (split_mnames _ ?t ?ls)] |- _ ]
- => is_var ls; rewrite (H ls) in H'
- | [ H : ?x = Some _, H' : ?x = None |- _ ] => congruence
- | [ H : ?x = Some ?a, H' : ?x = Some ?b |- _ ]
- => assert (a = b) by congruence; (subst a || subst b)
- end ].
- Qed.
-
- Lemma snd_split_mnames_skipn
- (t : flat_type base_type_code) (ls : list MName)
- : snd (split_mnames force t ls) = skipn (count_pairs t) ls.
- Proof using Type. rewrite split_mnames_firstn_skipn; reflexivity. Qed.
- Lemma fst_split_mnames_firstn
- (t : flat_type base_type_code) (ls : list MName)
- : fst (split_mnames force t ls) = fst (split_mnames force t (firstn (count_pairs t) ls)).
- Proof using Type. rewrite split_mnames_firstn_skipn at 1; reflexivity. Qed.
-
- Lemma mname_list_unique_firstn_skipn n ls
- : mname_list_unique force ls
- -> (mname_list_unique force (firstn n ls)
- /\ mname_list_unique force (skipn n ls)).
- Proof using Type.
- unfold mname_list_unique; intro H; split; intros k N;
- rewrite <- ?firstn_map, <- ?skipn_map, ?skipn_skipn, ?firstn_firstn_min, ?firstn_skipn_add;
- intros; eapply H; try eassumption.
- { apply Min.min_case_strong.
- { match goal with H : _ |- _ => rewrite skipn_firstn in H end;
- eauto using In_firstn. }
- { intro; match goal with H : _ |- _ => rewrite skipn_all in H by (rewrite firstn_length; omega * ) end.
- simpl in *; tauto. } }
- { eauto using In_skipn. }
- Qed.
- Definition mname_list_unique_firstn n ls
- : mname_list_unique force ls -> mname_list_unique force (firstn n ls)
- := fun H => proj1 (@mname_list_unique_firstn_skipn n ls H).
- Definition mname_list_unique_skipn n ls
- : mname_list_unique force ls -> mname_list_unique force (skipn n ls)
- := fun H => proj2 (@mname_list_unique_firstn_skipn n ls H).
- Lemma mname_list_unique_nil
- : mname_list_unique force nil.
- Proof using Type.
- unfold mname_list_unique; simpl; intros ??.
- rewrite firstn_nil, skipn_nil; simpl; auto.
- Qed.
- End monad.
-
- Lemma split_onames_firstn_skipn
- (t : flat_type base_type_code) (ls : list (option Name))
- : split_onames t ls
- = (fst (split_onames t (firstn (count_pairs t) ls)),
- skipn (count_pairs t) ls).
- Proof using Type. apply split_mnames_firstn_skipn. Qed.
- Lemma snd_split_onames_skipn
- (t : flat_type base_type_code) (ls : list (option Name))
- : snd (split_onames t ls) = skipn (count_pairs t) ls.
- Proof using Type. apply snd_split_mnames_skipn. Qed.
- Lemma fst_split_onames_firstn
- (t : flat_type base_type_code) (ls : list (option Name))
- : fst (split_onames t ls) = fst (split_onames t (firstn (count_pairs t) ls)).
- Proof using Type. apply fst_split_mnames_firstn. Qed.
-
- Lemma oname_list_unique_firstn n (ls : list (option Name))
- : oname_list_unique ls -> oname_list_unique (firstn n ls).
- Proof using Type. apply mname_list_unique_firstn. Qed.
- Lemma oname_list_unique_skipn n (ls : list (option Name))
- : oname_list_unique ls -> oname_list_unique (skipn n ls).
- Proof using Type. apply mname_list_unique_skipn. Qed.
- Lemma oname_list_unique_specialize (ls : list (option Name))
- : oname_list_unique ls
- -> forall k n,
- List.In (Some n) (firstn k ls)
- -> List.In (Some n) (skipn k ls)
- -> False.
- Proof using Type.
- intros H k n; specialize (H k n).
- rewrite map_id in H; assumption.
- Qed.
- Definition oname_list_unique_nil : oname_list_unique (@nil (option Name))
- := mname_list_unique_nil _ (fun x => x).
-
-
- Lemma split_names_firstn_skipn
- (t : flat_type base_type_code) (ls : list Name)
- : split_names t ls
- = (fst (split_names t (firstn (count_pairs t) ls)),
- skipn (count_pairs t) ls).
- Proof using Type. apply split_mnames_firstn_skipn. Qed.
- Lemma snd_split_names_skipn
- (t : flat_type base_type_code) (ls : list Name)
- : snd (split_names t ls) = skipn (count_pairs t) ls.
- Proof using Type. apply snd_split_mnames_skipn. Qed.
- Lemma fst_split_names_firstn
- (t : flat_type base_type_code) (ls : list Name)
- : fst (split_names t ls) = fst (split_names t (firstn (count_pairs t) ls)).
- Proof using Type. apply fst_split_mnames_firstn. Qed.
-
- Lemma name_list_unique_firstn n (ls : list Name)
- : name_list_unique ls -> name_list_unique (firstn n ls).
- Proof using Type.
- unfold name_list_unique; intro H; apply oname_list_unique_firstn with (n:=n) in H.
- rewrite <- firstn_map; assumption.
- Qed.
- Lemma name_list_unique_skipn n (ls : list Name)
- : name_list_unique ls -> name_list_unique (skipn n ls).
- Proof using Type.
- unfold name_list_unique; intro H; apply oname_list_unique_skipn with (n:=n) in H.
- rewrite <- skipn_map; assumption.
- Qed.
- Lemma name_list_unique_specialize (ls : list Name)
- : name_list_unique ls
- -> forall k n,
- List.In n (firstn k ls)
- -> List.In n (skipn k ls)
- -> False.
- Proof using Type.
- intros H k n; specialize (H k n).
- rewrite !map_id, !firstn_map, !skipn_map in H.
- eauto using in_map.
- Qed.
- Definition name_list_unique_nil : name_list_unique nil
- := mname_list_unique_nil _ (@Some Name).
-
- Lemma length_fst_split_names_Some_iff
- (t : flat_type base_type_code) (ls : list Name)
- : fst (split_names t ls) <> None <-> List.length ls >= count_pairs t.
- Proof using Type.
- revert ls; induction t; intros ls;
- try solve [ destruct ls; simpl; intuition (omega || congruence) ].
- repeat first [ progress simpl in *
- | progress break_innermost_match_step
- | progress specialize_by congruence
- | progress specialize_by omega
- | rewrite snd_split_names_skipn in *
- | progress intros
- | congruence
- | omega
- | match goal with
- | [ H : forall ls, fst (split_names ?t ls) <> None <-> _, H' : fst (split_names ?t ?ls') = _ |- _ ]
- => specialize (H ls'); rewrite H' in H
- | [ H : _ |- _ ] => rewrite skipn_length in H
- end
- | progress split_iff
- | match goal with
- | [ |- iff _ _ ] => split
- end ].
- Qed.
-
- Lemma length_fst_split_names_None_iff
- (t : flat_type base_type_code) (ls : list Name)
- : fst (split_names t ls) = None <-> List.length ls < count_pairs t.
- Proof using Type.
- destruct (length_fst_split_names_Some_iff t ls).
- destruct (le_lt_dec (count_pairs t) (List.length ls)); specialize_by omega;
- destruct (fst (split_names t ls)); split; try intuition (congruence || omega).
- Qed.
-
- Lemma split_onames_split_names (t : flat_type base_type_code) (ls : list Name)
- : split_onames t (List.map Some ls)
- = (fst (split_names t ls), List.map Some (snd (split_names t ls))).
- Proof using Type.
- revert ls; induction t;
- try solve [ destruct ls; reflexivity ].
- repeat first [ progress simpl in *
- | progress intros
- | rewrite snd_split_names_skipn
- | rewrite snd_split_onames_skipn
- | rewrite skipn_map
- | match goal with
- | [ H : forall ls, split_onames ?t (map Some ls) = _ |- context[split_onames ?t (map Some ?ls')] ]
- => specialize (H ls')
- end
- | break_innermost_match_step
- | progress inversion_prod
- | congruence ].
- Qed.
-End language.
diff --git a/src/Compilers/Named/PositiveContext.v b/src/Compilers/Named/PositiveContext.v
deleted file mode 100644
index 7b57098de..000000000
--- a/src/Compilers/Named/PositiveContext.v
+++ /dev/null
@@ -1,8 +0,0 @@
-Require Import Coq.FSets.FMapPositive.
-Require Import Crypto.Compilers.Named.FMapContext.
-
-Module PositiveContext := FMapContext PositiveMap.
-Notation PositiveContext := PositiveContext.FMapContext.
-Notation PositiveContext_nd := PositiveContext.FMapContext_nd.
-Definition PositiveContextOk := PositiveContext.FMapContextOk (fun x y pf => pf).
-Global Existing Instance PositiveContextOk.
diff --git a/src/Compilers/Named/PositiveContext/Defaults.v b/src/Compilers/Named/PositiveContext/Defaults.v
deleted file mode 100644
index 066de61cb..000000000
--- a/src/Compilers/Named/PositiveContext/Defaults.v
+++ /dev/null
@@ -1,29 +0,0 @@
-Require Import Coq.Lists.List.
-Require Import Coq.Numbers.BinNums.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.CountLets.
-Require Import Crypto.Compilers.CountLets.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Definition default_names_forf {var} (dummy : forall t, var t) {t} (e : exprf base_type_code op (var:=var) t) : list positive
- := map BinPos.Pos.of_succ_nat (seq 0 (count_let_bindersf dummy e)).
- Definition default_names_for {var} (dummy : forall t, var t) {t} (e : expr base_type_code op (var:=var) t) : list positive
- := map BinPos.Pos.of_succ_nat (seq 0 (count_binders dummy e)).
- Definition DefaultNamesFor {t} (e : Expr base_type_code op t) : list positive
- := map BinPos.Pos.of_succ_nat (seq 0 (CountBinders e)).
-End language.
-
-Module Named.
- Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Definition default_names_forf {Name} {t} (e : Named.exprf base_type_code op Name t) : list positive
- := map BinPos.Pos.of_succ_nat (seq 0 (Named.CountLets.count_let_bindersf e)).
- Definition default_names_for {Name} {t} (e : Named.expr base_type_code op Name t) : list positive
- := map BinPos.Pos.of_succ_nat (seq 0 (Named.CountLets.count_binders e)).
- End language.
-End Named.
diff --git a/src/Compilers/Named/PositiveContext/DefaultsProperties.v b/src/Compilers/Named/PositiveContext/DefaultsProperties.v
deleted file mode 100644
index 2b46de136..000000000
--- a/src/Compilers/Named/PositiveContext/DefaultsProperties.v
+++ /dev/null
@@ -1,38 +0,0 @@
-Require Import Coq.Lists.List.
-Require Import Coq.Numbers.BinNums.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.CountLets.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.NatUtil.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Lemma name_list_unique_map_pos_of_succ_nat_seq a b
- : name_list_unique (map BinPos.Pos.of_succ_nat (seq a b)).
- Proof using Type.
- unfold name_list_unique, oname_list_unique, mname_list_unique.
- intros k n.
- rewrite !map_map, firstn_map, skipn_map, firstn_seq, skipn_seq.
- rewrite !in_map_iff; intros; destruct_head' ex; destruct_head' and; inversion_option; subst.
- match goal with H : _ |- _ => apply Pnat.SuccNat2Pos.inj in H end; subst.
- rewrite in_seq in *.
- omega *.
- Qed.
-
- Lemma name_list_unique_default_names_forf {var dummy t e}
- : name_list_unique (@default_names_forf base_type_code op var dummy t e).
- Proof using Type. apply name_list_unique_map_pos_of_succ_nat_seq. Qed.
- Lemma name_list_unique_default_names_for {var dummy t e}
- : name_list_unique (@default_names_for base_type_code op var dummy t e).
- Proof using Type. apply name_list_unique_map_pos_of_succ_nat_seq. Qed.
- Lemma name_list_unique_DefaultNamesFor {t e}
- : name_list_unique (@DefaultNamesFor base_type_code op t e).
- Proof using Type. apply name_list_unique_map_pos_of_succ_nat_seq. Qed.
-End language.
diff --git a/src/Compilers/Named/RegisterAssign.v b/src/Compilers/Named/RegisterAssign.v
deleted file mode 100644
index 0b6c9b7b9..000000000
--- a/src/Compilers/Named/RegisterAssign.v
+++ /dev/null
@@ -1,89 +0,0 @@
-(** * Reassign registers *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.NameUtil.
-Require Import Crypto.Util.Decidable.
-
-Local Notation eta x := (fst x, snd x).
-
-Local Open Scope ctype_scope.
-Delimit Scope nexpr_scope with nexpr.
-Section language.
- Context (base_type_code : Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
-
- Section internal.
- Context (InName OutName : Type)
- {InContext : Context InName (fun _ : base_type_code => OutName)}
- {ReverseContext : Context OutName (fun _ : base_type_code => InName)}
- (InName_beq : InName -> InName -> bool).
-
- Definition register_reassignf_step
- (register_reassignf : forall (ctxi : InContext) (ctxr : ReverseContext)
- {t} (e : exprf InName t) (new_names : list (option OutName)),
- option (exprf OutName t))
- (ctxi : InContext) (ctxr : ReverseContext)
- {t} (e : exprf InName t) (new_names : list (option OutName))
- : option (exprf OutName t)
- := match e in Named.exprf _ _ _ t return option (exprf _ t) with
- | TT => Some TT
- | Var t' name => match lookupb t' ctxi name with
- | Some new_name
- => match lookupb t' ctxr new_name with
- | Some name'
- => if InName_beq name name'
- then Some (Var new_name)
- else None
- | None => None
- end
- | None => None
- end
- | Op _ _ op args
- => option_map (Op op) (@register_reassignf ctxi ctxr _ args new_names)
- | LetIn tx n ex _ eC
- => let '(n', new_names') := eta (split_onames tx new_names) in
- match n', @register_reassignf ctxi ctxr _ ex nil with
- | Some n', Some x
- => let ctxi := extend ctxi n n' in
- let ctxr := extend ctxr n' n in
- option_map (LetIn tx n' x) (@register_reassignf ctxi ctxr _ eC new_names')
- | _, _
- => let ctxi := remove ctxi n in
- @register_reassignf ctxi ctxr _ eC new_names'
- end
- | Pair _ ex _ ey
- => match @register_reassignf ctxi ctxr _ ex nil, @register_reassignf ctxi ctxr _ ey nil with
- | Some x, Some y
- => Some (Pair x y)
- | _, _ => None
- end
- end.
- Fixpoint register_reassignf ctxi ctxr {t} e new_names
- := @register_reassignf_step (@register_reassignf) ctxi ctxr t e new_names.
-
- Definition register_reassign (ctxi : InContext) (ctxr : ReverseContext)
- {t} (e : expr InName t) (new_names : list (option OutName))
- : option (expr OutName t)
- := match e in Named.expr _ _ _ t return option (expr _ t) with
- | Abs src _ n f
- => let '(n', new_names') := eta (split_onames src new_names) in
- match n' with
- | Some n'
- => let ctxi := extend (t:=src) ctxi n n' in
- let ctxr := extend (t:=src) ctxr n' n in
- option_map (Abs n') (register_reassignf ctxi ctxr f new_names')
- | None => None
- end
- end.
- End internal.
-End language.
-
-Global Arguments register_reassign {_ _ _ _ _ _} _ ctxi ctxr {t} e _.
-Global Arguments register_reassignf {_ _ _ _ _ _} _ ctxi ctxr {t} e _.
diff --git a/src/Compilers/Named/RegisterAssignInterp.v b/src/Compilers/Named/RegisterAssignInterp.v
deleted file mode 100644
index be0e9cace..000000000
--- a/src/Compilers/Named/RegisterAssignInterp.v
+++ /dev/null
@@ -1,239 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.RegisterAssign.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.HProp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Local Open Scope nexpr_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- {interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {InName OutName : Type}
- {InContext : Context InName (fun _ : base_type_code => OutName)}
- {ReverseContext : Context OutName (fun _ : base_type_code => InName)}
- {InContextOk : ContextOk InContext}
- {ReverseContextOk : ContextOk ReverseContext}
- {InName_beq : InName -> InName -> bool}
- {InterpInContext : Context InName interp_base_type}
- {InterpOutContext : Context OutName interp_base_type}
- {InterpInContextOk : ContextOk InterpInContext}
- {InterpOutContextOk : ContextOk InterpOutContext}
- {base_type_code_dec : DecidableRel (@eq base_type_code)}
- {OutName_dec : DecidableRel (@eq OutName)}
- (InName_bl : forall n1 n2, InName_beq n1 n2 = true -> n1 = n2)
- (InName_lb : forall n1 n2, n1 = n2 -> InName_beq n1 n2 = true).
-
- Local Instance InName_dec : DecidableRel (@eq InName)
- := dec_rel_of_bool_dec_rel InName_beq InName_bl InName_lb.
-
- Local Notation register_reassignf := (@register_reassignf base_type_code op InName OutName InContext ReverseContext InName_beq).
- Local Notation register_reassign := (@register_reassign base_type_code op InName OutName InContext ReverseContext InName_beq).
-
- Local Ltac t_find T :=
- induction T;
- repeat first [ progress subst
- | progress inversion_option
- | progress simpl in *
- | break_innermost_match_step
- | break_innermost_match_hyps_step
- | progress destruct_head'_and
- | repeat apply conj; congruence
- | progress cbv [cast_if_eq] in *
- | progress eliminate_hprop_eq
- | solve [ eauto ]
- | match goal with
- | [ |- context[@find_Name_and_val ?base_type_code ?Name ?base_type_code_dec ?Name_dec ?var' ?t ?n ?T ?N ?V ?default] ]
- => lazymatch default with
- | None => fail
- | _ => rewrite (@find_Name_and_val_split base_type_code base_type_code_dec Name Name_dec var' t n T N V default)
- end
- | [ H : context[@find_Name_and_val ?base_type_code ?Name ?base_type_code_dec ?Name_dec ?var' ?t ?n ?T ?N ?V ?default] |- _ ]
- => lazymatch default with
- | None => fail
- | _ => rewrite (@find_Name_and_val_split base_type_code base_type_code_dec Name Name_dec var' t n T N V default) in H
- end
- | [ H : forall a b, find_Name_and_val _ _ _ _ _ _ _ = Some _ -> _ |- find_Name_and_val _ _ _ _ _ _ _ = Some _ -> _ ]
- => let H' := fresh in intro H'; pose proof (H _ _ H')
- | [ H : forall a b, find_Name_and_val _ _ _ _ _ _ _ = Some _ -> _, H' : find_Name_and_val _ _ _ _ _ _ _ = Some _ |- _ ]
- => specialize (H _ _ H')
- | [ H : ?x = Some ?a, H' : ?y = Some ?b |- _ ]
- => assert (a = b) by congruence; progress subst
- end ].
-
-
- Lemma find_Name_and_val_OutToIn t T NI NO n_in n_out
- : find_Name_and_val base_type_code_dec OutName_dec (T:=T) t n_out NO NI None = Some n_in
- -> find_Name InName_dec n_in NI <> None.
- Proof using Type. t_find T. Qed.
- Lemma find_Name_and_val_InToOut t T NI NO n_in n_out
- : find_Name_and_val base_type_code_dec InName_dec (T:=T) t n_in NI NO None = Some n_out
- -> find_Name OutName_dec n_out NO <> None.
- Proof using Type. t_find T. Qed.
-
- Lemma lookupb_extend_helper {ctxi : InContext} {ctxr : ReverseContext} {t T NI NO n_in n_out}
- (H0 : lookupb t (extend (t:=T) ctxi NI NO) n_in = Some n_out)
- (H1 : lookupb t (extend (t:=T) ctxr NO NI) n_out = Some n_in)
- : ((lookupb t ctxi n_in = Some n_out /\ lookupb t ctxr n_out = Some n_in)
- /\ (find_Name _ n_in NI = None /\ find_Name _ n_out NO = None))
- \/ (find_Name_and_val _ _ t n_in NI NO None = Some n_out
- /\ find_Name_and_val _ _ t n_out NO NI None = Some n_in).
- Proof using InContextOk ReverseContextOk.
- rewrite (@lookupb_extend base_type_code _ InName _) in H0 by assumption.
- rewrite (@lookupb_extend base_type_code _ OutName _) in H1 by assumption.
- repeat first [ progress subst
- | match goal with
- | [ H : find_Name_and_val _ _ _ _ _ _ ?x = _ |- _ ]
- => lazymatch x with
- | None => fail
- | _ => rewrite find_Name_and_val_split in H
- end
- end
- | break_innermost_match_hyps_step
- | congruence
- | solve [ auto
- | exfalso; eapply find_Name_and_val_OutToIn; eassumption
- | exfalso; eapply find_Name_and_val_InToOut; eassumption ] ].
- Qed.
-
- Lemma find_Name_and_val_same_val {var' t T n_in n_out NI NO V v1 v2}
- (H0 : find_Name_and_val base_type_code_dec InName_dec t (T:=T) (var':=var') n_in NI V None = Some v1)
- (H1 : find_Name_and_val base_type_code_dec OutName_dec t (T:=T) n_out NO V None = Some v2)
- (H2 : find_Name_and_val base_type_code_dec InName_dec t n_in NI NO None = Some n_out)
- (H3 : find_Name_and_val base_type_code_dec OutName_dec t n_out NO NI None = Some n_in)
- : v1 = v2.
- Proof using Type.
- t_find T;
- match goal with
- | [ H : _ = None |- _ ]
- => first [ eapply find_Name_and_val_OutToIn in H; [ | eassumption ]
- | eapply find_Name_and_val_InToOut in H; [ | eassumption ] ];
- destruct H
- end.
- Qed.
-
- Lemma find_Name_and_val_same_oval {var' t T n_in n_out NI NO V}
- (H2 : find_Name_and_val base_type_code_dec InName_dec t n_in NI NO None = Some n_out)
- (H3 : find_Name_and_val base_type_code_dec OutName_dec t n_out NO NI None = Some n_in)
- : find_Name_and_val base_type_code_dec InName_dec t (T:=T) (var':=var') n_in NI V None
- = find_Name_and_val base_type_code_dec OutName_dec t (T:=T) n_out NO V None.
- Proof using Type.
- t_find T;
- match goal with
- | [ H : _ = None |- _ ]
- => first [ eapply find_Name_and_val_OutToIn in H; [ | eassumption ]
- | eapply find_Name_and_val_InToOut in H; [ | eassumption ] ];
- destruct H
- end.
- Qed.
-
- Local Ltac t :=
- repeat first [ reflexivity
- | assumption
- | progress subst
- | progress inversion_option
- | progress simpl in *
- | progress intros *
- | progress intros
- | progress destruct_head'_and
- | progress destruct_head'_or
- | match goal with
- | [ H : InName_beq _ _ = true |- _ ] => apply InName_bl in H
- | [ H : _ = Some ?v1, H' : _ = Some ?v2 |- _ ]
- => is_var v1; is_var v2;
- assert (v1 = v2) by eauto; progress subst
- | [ H : lookupb (extend _ _ _) _ = Some _, H' : lookupb (extend _ _ _) _ = Some _ |- _ ]
- => pose proof (lookupb_extend_helper H H'); clear H H'
- | [ H : find_Name_and_val _ _ _ _ _ _ ?x = _ |- _ ]
- => lazymatch x with
- | None => fail
- | _ => rewrite find_Name_and_val_split in H
- end
- | [ |- context[@find_Name_and_val ?base_type_code ?Name ?base_type_code_dec ?Name_dec ?var' ?t ?n ?T ?N ?V ?default] ]
- => lazymatch default with
- | None => fail
- | _ => rewrite (@find_Name_and_val_split base_type_code base_type_code_dec Name Name_dec var' t n T N V default)
- end
- | [ H : _ |- _ ]
- => first [ rewrite (@lookupb_remove base_type_code InName _) in H
- | rewrite (@lookupb_remove base_type_code OutName _) in H
- | rewrite (@lookupb_extend base_type_code _ InName _) in H
- | rewrite (@lookupb_extend base_type_code _ OutName _) in H
- | rewrite find_Name_and_val_different in H by assumption ]
- end
- | break_innermost_match_step
- | break_innermost_match_hyps_step
- | progress cbv [option_map LetIn.Let_In] in *
- | solve [ eauto using find_Name_and_val_same_val ]
- | match goal with
- | [ H : _ |- ?v1 = ?v2 ]
- => is_var v1; is_var v2;
- eapply H; [ | eassumption | eassumption | eassumption ]; clear H
- end ].
-
- Lemma interpf_register_reassignf
- ctxi ctxr t e new_names
- (ctxi_interp : InterpInContext)
- (ctxr_interp : InterpOutContext)
- eout
- v1 v2
- : (forall t (n_in : InName) (n_out : OutName) v1 v2,
- lookupb t ctxi n_in = Some n_out
- -> lookupb t ctxr n_out = Some n_in
- -> lookupb t ctxi_interp n_in = Some v1
- -> lookupb t ctxr_interp n_out = Some v2
- -> v1 = v2)
- -> @register_reassignf ctxi ctxr t e new_names = Some eout
- -> interpf (interp_op:=interp_op) (ctx:=ctxr_interp) eout = Some v1
- -> interpf (interp_op:=interp_op) (ctx:=ctxi_interp) e = Some v2
- -> v1 = v2.
- Proof using InContextOk InName_bl InName_lb InterpInContextOk InterpOutContextOk OutName_dec ReverseContextOk base_type_code_dec.
- revert ctxi ctxr new_names ctxi_interp ctxr_interp eout v1 v2.
- induction e; t.
- Qed.
-
- Lemma interp_register_reassign
- ctxi ctxr t e new_names
- (ctxi_interp : InterpInContext)
- (ctxr_interp : InterpOutContext)
- eout
- v1 v2 x
- : (forall t (n_in : InName) (n_out : OutName) v1 v2,
- lookupb t ctxi n_in = Some n_out
- -> lookupb t ctxr n_out = Some n_in
- -> lookupb t ctxi_interp n_in = Some v1
- -> lookupb t ctxr_interp n_out = Some v2
- -> v1 = v2)
- -> @register_reassign ctxi ctxr t e new_names = Some eout
- -> interp (interp_op:=interp_op) (ctx:=ctxr_interp) eout x = Some v1
- -> interp (interp_op:=interp_op) (ctx:=ctxi_interp) e x = Some v2
- -> v1 = v2.
- Proof using InContextOk InName_bl InName_lb InterpInContextOk InterpOutContextOk OutName_dec ReverseContextOk base_type_code_dec.
- destruct e; unfold interp, register_reassign, option_map in *.
- break_innermost_match; try congruence.
- intros; inversion_option; subst; simpl in *.
- eapply interpf_register_reassignf; [ | eassumption.. ].
- t.
- Qed.
-
- Lemma Interp_register_reassign
- t e new_names
- eout
- v1 v2 x
- : @register_reassign empty empty t e new_names = Some eout
- -> Interp (Context:=InterpOutContext) (interp_op:=interp_op) eout x = Some v1
- -> Interp (Context:=InterpInContext) (interp_op:=interp_op) e x = Some v2
- -> v1 = v2.
- Proof using InContextOk InName_bl InName_lb InterpInContextOk InterpOutContextOk OutName_dec ReverseContextOk base_type_code_dec.
- apply interp_register_reassign; intros *.
- rewrite !lookupb_empty; congruence.
- Qed.
-End language.
diff --git a/src/Compilers/Named/SmartMap.v b/src/Compilers/Named/SmartMap.v
deleted file mode 100644
index 76b2eec58..000000000
--- a/src/Compilers/Named/SmartMap.v
+++ /dev/null
@@ -1,20 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Named.Syntax.
-
-Module Export Named.
- Section language.
- Context {base_type_code : Type}
- {interp_base_type : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {Name : Type}.
-
- (** [SmartVar] is like [Var], except that it inserts
- pair-projections and [Pair] as necessary to handle
- [flat_type], and not just [base_type_code] *)
- Definition SmartVar {t} : interp_flat_type (fun _ => Name) t -> @exprf base_type_code op Name t
- := smart_interp_flat_map (f:=fun _ => Name) (g:=@exprf _ _ _) (fun t => Var) TT (fun A B x y => Pair x y).
- End language.
-End Named.
-
-Global Arguments SmartVar {_ _ _ _} _.
diff --git a/src/Compilers/Named/Syntax.v b/src/Compilers/Named/Syntax.v
deleted file mode 100644
index 4cc4ccbf2..000000000
--- a/src/Compilers/Named/Syntax.v
+++ /dev/null
@@ -1,92 +0,0 @@
-(** * Named Representation of Gallina *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.LetIn.
-
-Local Open Scope ctype_scope.
-Local Open Scope expr_scope.
-Delimit Scope nexpr_scope with nexpr.
-Module Export Named.
- Section language.
- Context (base_type_code : Type)
- (interp_base_type : base_type_code -> Type)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (Name : Type).
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type_gen := interp_flat_type.
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
-
- Inductive exprf : flat_type -> Type :=
- | TT : exprf Unit
- | Var {t : base_type_code} : Name -> exprf (Tbase t)
- | Op {t1 tR} : op t1 tR -> exprf t1 -> exprf tR
- | LetIn : forall {tx}, interp_flat_type_gen (fun _ => Name) tx -> exprf tx -> forall {tC}, exprf tC -> exprf tC
- | Pair : forall {t1}, exprf t1 -> forall {t2}, exprf t2 -> exprf (Prod t1 t2).
- Bind Scope nexpr_scope with exprf.
- Inductive expr : type -> Type :=
- | Abs {src dst} : interp_flat_type_gen (fun _ => Name) src -> exprf dst -> expr (Arrow src dst).
- Bind Scope nexpr_scope with expr.
- Definition Abs_name {t} (e : expr t) : interp_flat_type_gen (fun _ => Name) (domain t)
- := match e with Abs _ _ n f => n end.
- Definition invert_Abs {t} (e : expr t) : exprf (codomain t)
- := match e with Abs _ _ n f => f end.
-
- Section with_val_context.
- Context (Context : Context Name interp_base_type)
- (interp_op : forall src dst, op src dst -> interp_flat_type src -> interp_flat_type dst).
-
- Fixpoint interpf
- {t} (ctx : Context) (e : exprf t)
- : option (interp_flat_type t)
- := match e in exprf t return option (interp_flat_type t) with
- | Var t' x => lookupb t' ctx x
- | TT => Some tt
- | Pair _ ex _ ey
- => match @interpf _ ctx ex, @interpf _ ctx ey with
- | Some xv, Some yv => Some (xv, yv)%core
- | None, _ | _, None => None
- end
- | Op _ _ opc args
- => option_map (@interp_op _ _ opc) (@interpf _ ctx args)
- | LetIn _ n ex _ eC
- => match @interpf _ ctx ex with
- | Some xv
- => dlet x := xv in
- @interpf _ (extend ctx n x) eC
- | None => None
- end
- end.
-
- Definition interp {t} (ctx : Context) (e : expr t)
- : interp_flat_type (domain t) -> option (interp_flat_type (codomain t))
- := fun v => @interpf _ (extend ctx (Abs_name e) v) (invert_Abs e).
-
- Definition Interp {t} (e : expr t)
- : interp_flat_type (domain t) -> option (interp_flat_type (codomain t))
- := interp empty e.
- End with_val_context.
- End language.
-End Named.
-
-Global Arguments TT {_ _ _}.
-Global Arguments Var {_ _ _ _} _.
-Global Arguments Op {_ _ _ _ _} _ _.
-Global Arguments LetIn {_ _ _} _ _ _ {_} _.
-Global Arguments Pair {_ _ _ _} _ {_} _.
-Global Arguments Abs {_ _ _ _ _} _ _.
-Global Arguments invert_Abs {_ _ _ _} _.
-Global Arguments Abs_name {_ _ _ _} _.
-Global Arguments interpf {_ _ _ _ _ interp_op t ctx} _.
-Global Arguments interp {_ _ _ _ _ interp_op t ctx} _ _.
-Global Arguments Interp {_ _ _ _ _ interp_op t} _ _.
-
-Notation "'nlet' x := A 'in' b" := (LetIn _ x A%nexpr b%nexpr) : nexpr_scope.
-Notation "'nlet' x : tx := A 'in' b" := (LetIn tx%ctype x%core A%nexpr b%nexpr) : nexpr_scope.
-Notation "'λn' x .. y , t" := (Abs x .. (Abs y t%nexpr) .. ) : nexpr_scope.
-Notation "( x , y , .. , z )" := (Pair .. (Pair x%nexpr y%nexpr) .. z%nexpr) : nexpr_scope.
-Notation "()" := TT : nexpr_scope.
diff --git a/src/Compilers/Named/WeakListContext.v b/src/Compilers/Named/WeakListContext.v
deleted file mode 100644
index f1c4a46e3..000000000
--- a/src/Compilers/Named/WeakListContext.v
+++ /dev/null
@@ -1,10 +0,0 @@
-(** * Context made from an associative list *)
-Require Import Coq.FSets.FMapWeakList.
-Require Import Coq.FSets.FMapInterface.
-Require Import Crypto.Compilers.Named.FMapContext.
-
-Module WeakListContext (E : DecidableType).
- Module WL := FMapWeakList.Make E.
- Module Context := FMapContext WL.
- Include Context.
-End WeakListContext.
diff --git a/src/Compilers/Named/Wf.v b/src/Compilers/Named/Wf.v
deleted file mode 100644
index 79f620edb..000000000
--- a/src/Compilers/Named/Wf.v
+++ /dev/null
@@ -1,56 +0,0 @@
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Util.PointedProp.
-
-Module Export Named.
- Section language.
- Context {base_type_code Name : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Section with_var.
- Context {var}
- (Context : @Context base_type_code Name var).
-
- Section gen.
- Context (quantify : forall tx, (interp_flat_type var tx -> option pointed_Prop) -> option pointed_Prop).
-
- Fixpoint wff_gen (ctx : Context) {t} (e : @exprf base_type_code op Name t) : option pointed_Prop
- := match e with
- | TT => Some trivial
- | Var t n => match lookupb t ctx n return bool with
- | Some _ => true
- | None => false
- end
- | Op _ _ op args => @wff_gen ctx _ args
- | LetIn _ n ex _ eC => @wff_gen ctx _ ex
- /\ quantify _ (fun v => @wff_gen (extend ctx n v) _ eC)
- | Pair _ ex _ ey => @wff_gen ctx _ ex /\ @wff_gen ctx _ ey
- end%option_pointed_prop.
- End gen.
-
- Definition wff := wff_gen (fun tx f => inject (forall v, prop_of_option (f v))).
- Definition wf (ctx : Context) {t} (e : @expr base_type_code op Name t) : Prop
- := forall v, prop_of_option (@wff (extend ctx (Abs_name e) v) _ (invert_Abs e)).
- End with_var.
-
- Section unit.
- Context (Context : @Context base_type_code Name (fun _ => unit)).
-
- Local Notation TT := (SmartValf _ (fun _ => tt) _).
- Definition wff_unit := wff_gen Context (fun tx f => f TT).
- Definition wf_unit ctx {t} (e : @expr base_type_code op Name t) : option pointed_Prop
- := @wff_unit (extend ctx (Abs_name e) TT) _ (invert_Abs e).
- End unit.
-
- Definition Wf (Context : forall var, @Context base_type_code Name var) {t} (e : @expr base_type_code op Name t)
- := forall var, wf (Context var) empty e.
- End language.
-End Named.
-
-Global Arguments wff_gen {_ _ _ _ _} quantify ctx {t} _.
-Global Arguments wff_unit {_ _ _ _} ctx {t} _.
-Global Arguments wf_unit {_ _ _ _} ctx {t} _.
-Global Arguments wff {_ _ _ _ _} ctx {t} _.
-Global Arguments wf {_ _ _ _ _} ctx {t} _.
-Global Arguments Wf {_ _ _} Context {t} _.
diff --git a/src/Compilers/Named/WfFromUnit.v b/src/Compilers/Named/WfFromUnit.v
deleted file mode 100644
index e00a79274..000000000
--- a/src/Compilers/Named/WfFromUnit.v
+++ /dev/null
@@ -1,82 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextProperties.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Option.
-
-Section language.
- Context {base_type_code Name : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {base_type_code_dec : DecidableRel (@eq base_type_code)}
- {Name_dec : DecidableRel (@eq Name)}.
-
- Section with_var.
- Context {var}
- (uContext : @Context base_type_code Name (fun _ => unit))
- (uContextOk : ContextOk uContext)
- (vContext : @Context base_type_code Name var)
- (vContextOk : ContextOk vContext).
-
- Local Ltac t :=
- repeat first [ progress simpl in *
- | progress intros
- | progress subst
- | progress inversion_option
- | congruence
- | tauto
- | solve [ eauto ]
- | break_innermost_match_step
- | break_innermost_match_hyps_step
- | progress destruct_head'_and
- | progress autorewrite with push_prop_of_option push_eq_Some_trivial in *
- | rewrite !(@lookupb_extend base_type_code _ Name _) by auto
- | rewrite (@find_Name_and_val_split base_type_code _ Name _) with (default := lookupb _ _)
- | match goal with
- | [ H : ?x = Some _ |- _ ]
- => assert (x = None) by (split_iff; eauto); congruence
- | [ |- _ /\ _ ] => split
- | [ H : _ |- prop_of_option _ ] => eapply H; [ | eassumption ]; clear H
- | [ |- context[find_Name_and_val ?tdec ?ndec ?b _ _ _ _ = None] ]
- => rewrite <- !(@find_Name_and_val_None_iff _ tdec _ ndec _ b)
- end ].
-
- Lemma wff_from_unit
- (vctx : vContext)
- (uctx : uContext)
- (Hctx : forall t n, lookupb t vctx n = None <-> lookupb t uctx n = None)
- {t} (e : @exprf base_type_code op Name t)
- : wff_unit uctx e = Some trivial -> prop_of_option (wff vctx e).
- Proof using Name_dec base_type_code_dec uContextOk vContextOk.
- revert uctx vctx Hctx; induction e; t.
- Qed.
-
- Lemma wf_from_unit
- (vctx : vContext)
- (uctx : uContext)
- (Hctx : forall t n, lookupb t vctx n = None <-> lookupb t uctx n = None)
- {t} (e : @expr base_type_code op Name t)
- : wf_unit uctx e = Some trivial -> wf vctx e.
- Proof using Name_dec base_type_code_dec uContextOk vContextOk.
- intros H ?; revert H; apply wff_from_unit; t.
- Qed.
- End with_var.
-
- Lemma Wf_from_unit
- (Context : forall var, @Context base_type_code Name var)
- (ContextOk : forall var, ContextOk (Context var))
- {t} (e : @expr base_type_code op Name t)
- : wf_unit (Context:=Context _) empty e = Some trivial -> Wf Context e.
- Proof using Name_dec base_type_code_dec.
- intros H ?; revert H; apply wf_from_unit; auto; intros.
- rewrite !lookupb_empty by auto; tauto.
- Qed.
-End language.
-
-Hint Resolve wf_from_unit Wf_from_unit wff_from_unit : wf.
diff --git a/src/Compilers/Named/WfInterp.v b/src/Compilers/Named/WfInterp.v
deleted file mode 100644
index e9cd8737c..000000000
--- a/src/Compilers/Named/WfInterp.v
+++ /dev/null
@@ -1,41 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Tactics.DestructHead.
-
-Section language.
- Context {base_type_code Name interp_base_type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}
- {Context : @Context base_type_code Name interp_base_type}.
-
- Lemma wff_interpf_not_None {ctx : Context} {t} {e : @exprf base_type_code op Name t}
- (Hwf : prop_of_option (wff ctx e))
- : @interpf base_type_code interp_base_type op Name Context interp_op t ctx e <> None.
- Proof using Type.
- revert dependent ctx; induction e;
- repeat first [ progress intros
- | progress simpl in *
- | progress unfold option_map, LetIn.Let_In in *
- | congruence
- | progress specialize_by_assumption
- | progress destruct_head' and
- | progress break_innermost_match_step
- | progress break_match_hyps
- | progress autorewrite with push_prop_of_option in *
- | progress specialize_by auto
- | solve [ intuition eauto ] ].
- Qed.
-
- Lemma wf_interp_not_None {ctx : Context} {t} {e : @expr base_type_code op Name t}
- (Hwf : wf ctx e)
- v
- : @interp base_type_code interp_base_type op Name Context interp_op t ctx e v <> None.
- Proof using Type.
- destruct e; unfold interp, wf in *; apply wff_interpf_not_None; auto.
- Qed.
-End language.
diff --git a/src/Compilers/Reify.v b/src/Compilers/Reify.v
deleted file mode 100644
index 4454f2af2..000000000
--- a/src/Compilers/Reify.v
+++ /dev/null
@@ -1,591 +0,0 @@
-(** * Exact reification of PHOAS Representation of Gallina *)
-(** The reification procedure goes through [InputSyntax], which allows
- judgmental equality of the denotation of the reified term. *)
-Require Import Coq.Strings.String.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Compilers.InputSyntax.
-Require Crypto.Compilers.Tuple.
-Require Import Crypto.Util.Tuple.
-Require Import Crypto.Util.Tactics.DebugPrint.
-Require Import Crypto.Util.SideConditions.CorePackages.
-(*Require Import Crypto.Util.Tactics.PrintContext.*)
-Require Import Crypto.Util.Tactics.Head.
-Require Import Crypto.Util.Tactics.SubstLet.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.Tactics.TransparentAssert.
-
-(** Change this with [Ltac reify_debug_level ::= constr:(1).] to get
- more debugging. *)
-Ltac reify_debug_level := constr:(0).
-Module Import ReifyDebugNotations.
- Export Compilers.Syntax.Notations.
- Export Util.LetIn.
- Open Scope string_scope.
-End ReifyDebugNotations.
-
-Tactic Notation "debug_enter_reify_idtac" ident(funname) uconstr(e)
- := idtac funname ": Attempting to reify:" e.
-Tactic Notation "debug_leave_reify_success_idtac" ident(funname) uconstr(e) uconstr(ret)
- := idtac funname ": Success in reifying:" e "as" ret.
-Tactic Notation "debug_leave_reify_failure_idtac" ident(funname) uconstr(e)
- := idtac funname ": Failure in reifying:" e.
-Ltac check_debug_level_then_Set _ :=
- let lvl := reify_debug_level in
- lazymatch type of lvl with
- | nat => constr:(Set)
- | ?T => constr_run_tac ltac:(fun _ => idtac "Error: reify_debug_level should have type nat but instead has type" T)
- end.
-Ltac debug0 tac :=
- constr_run_tac tac.
-Ltac debug1 tac :=
- let lvl := reify_debug_level in
- lazymatch lvl with
- | S _ => constr_run_tac tac
- | _ => check_debug_level_then_Set ()
- end.
-Ltac debug2 tac :=
- let lvl := reify_debug_level in
- lazymatch lvl with
- | S (S _) => constr_run_tac tac
- | _ => check_debug_level_then_Set ()
- end.
-Ltac debug3 tac :=
- let lvl := reify_debug_level in
- lazymatch lvl with
- | S (S (S _)) => constr_run_tac tac
- | _ => check_debug_level_then_Set ()
- end.
-Ltac debug_enter_reify_flat_type e := debug2 ltac:(fun _ => debug_enter_reify_idtac reify_flat_type e).
-Ltac debug_enter_reify_type e := debug2 ltac:(fun _ => debug_enter_reify_idtac reify_type e).
-Ltac debug_enter_reifyf e := debug2 ltac:(fun _ => debug_enter_reify_idtac reifyf e).
-Ltac debug_leave_reifyf_success e ret := debug3 ltac:(fun _ => debug_leave_reify_success_idtac reifyf e ret).
-Ltac debug_leave_reifyf_failure e
- := let dummy := debug0 ltac:(fun _ => debug_leave_reify_failure_idtac reifyf e) in
- constr:(I : I).
-Ltac debug_leave_reify_flat_type_failure e
- := let dummy := debug0 ltac:(fun _ => debug_leave_reify_failure_idtac reify_flat_type e) in
- constr:(I : I).
-Tactic Notation "idtac_reifyf_case" ident(case) :=
- idtac "reifyf:" case.
-Ltac debug_reifyf_case tac :=
- debug3 tac.
-Ltac debug_enter_reify_abs e := debug2 ltac:(fun _ => debug_enter_reify_idtac reify_abs e).
-
-Ltac refresh x :=
- (* Work around Coq 8.5 and 8.6 bug *)
- (* <https://coq.inria.fr/bugs/show_bug.cgi?id=4998> *)
- (* Avoid re-binding the Gallina variable referenced by Ltac [x] *)
- (* even if its Gallina name matches a Ltac in this tactic. *)
- let maybe_x := fresh x in
- let maybe_x := fresh x in
- let maybe_x := fresh x in
- let not_x := fresh x in
- not_x.
-Class reify_internal {varT} (var : varT) {eT} (e : eT) {T : Type} := Build_reify_internal : T.
-Class reify {varT} (var : varT) {eT} (e : eT) {T : Type} := Build_reify : T.
-Typeclasses Opaque reify_internal reify.
-Definition reify_var_for_in_is base_type_code {T} (x : T) (t : flat_type base_type_code) {eT} (e : eT) := False.
-Arguments reify_var_for_in_is _ {T} _ _ {eT} _.
-
-(** [reify] assumes that operations can be reified via the [reify_op]
- typeclass, which gets passed the type family of operations, the
- expression which is headed by an operation, and expects resolution
- to fill in a number of arguments (which [reifyf] will
- automatically curry), as well as the reified operator.
-
- We also assume that types can be reified via the [reify] typeclass
- with arguments [reify type <type to be reified>]. *)
-Class reify_op {opTF} (op_family : opTF) {opExprT} (opExpr : opExprT) (nargs : nat) {opT} (reified_op : opT)
- := Build_reify_op : True.
-Ltac strip_type_cast term := lazymatch term with ?term' => term' end.
-(** Override this to get a faster [reify_type] *)
-Ltac base_reify_type T :=
- strip_type_cast (_ : reify type T).
-Ltac reify_base_type T := base_reify_type T.
-Ltac reify_flat_type T :=
- let dummy := debug_enter_reify_flat_type T in
- match goal with
- | _
- => lazymatch T with
- | prod ?A ?B
- => let a := reify_flat_type A in
- let b := reify_flat_type B in
- constr:(@Prod _ a b)
- | Syntax.interp_type _ (Tflat ?T)
- => T
- | Syntax.interp_flat_type _ ?T
- => T
- | Syntax.tuple _ _
- => T
- | Syntax.tuple' _ _
- => T
- | Tuple.tuple ?A ?n
- => let a := reify_flat_type A in
- constr:(@Syntax.tuple _ a n)
- | Tuple.tuple' ?A ?n
- => let a := reify_flat_type A in
- constr:(@Syntax.tuple' _ a n)
- | unit
- => Unit
- | _
- => let v := reify_base_type T in
- constr:(@Tbase _ v)
- end
- | _ => debug_leave_reify_flat_type_failure T
- end.
-Ltac reify_input_type T :=
- let dummy := debug_enter_reify_type T in
- lazymatch T with
- | (?A -> ?B)%type
- => let a := reify_flat_type A in
- let b := reify_input_type B in
- constr:(@Arrow _ a b)
- | InputSyntax.interp_type _ ?T
- => T
- end.
-Ltac reify_type T :=
- let dummy := debug_enter_reify_type T in
- lazymatch T with
- | (?A -> ?B)%type
- => let a := reify_flat_type A in
- let b := reify_flat_type B in
- constr:(@Syntax.Arrow _ a b)
- | Syntax.interp_type _ ?T
- => T
- end.
-
-Ltac reifyf_var x mkVar :=
- lazymatch goal with
- | _ : reify_var_for_in_is _ x ?t ?v |- _ => mkVar t v
- | _ => lazymatch x with
- | fst ?x' => reifyf_var x' ltac:(fun t v => lazymatch t with
- | Prod ?A ?B => mkVar A (fst v)
- end)
- | snd ?x' => reifyf_var x' ltac:(fun t v => lazymatch t with
- | Prod ?A ?B => mkVar B (snd v)
- end)
- end
- end.
-
-Inductive reify_result_helper :=
-| finished_value {T} (res : T)
-| context_value {TF} (resF : TF) {argT} (arg : argT)
-| op_info {T} (res : T)
-| reification_unsuccessful.
-
-(** Override this to get a faster [reify_op] *)
-Ltac base_reify_op op op_head expr :=
- let r := constr:(_ : reify_op op op_head _ _) in
- type of r.
-Ltac reify_op op op_head expr :=
- let t := base_reify_op op op_head expr in
- constr:(op_info t).
-
-Ltac debug_enter_reify_rec :=
- let lvl := reify_debug_level in
- match lvl with
- | S _ => idtac_goal
- | _ => idtac
- end.
-Ltac debug_leave_reify_rec e :=
- let lvl := reify_debug_level in
- match lvl with
- | S _ => idtac "<infomsg>reifyf success:" e "</infomsg>"
- | _ => idtac
- end.
-
-Ltac reifyf base_type_code interp_base_type op var e :=
- let reify_rec e := reifyf base_type_code interp_base_type op var e in
- let mkLetIn ex eC := constr:(LetIn (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) ex eC) in
- let mkPair ex ey := constr:(Pair (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) ex ey) in
- let mkVar T ex := constr:(Var (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) (t:=T) ex) in
- let mkConst T ex := constr:(Const (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) (t:=T) ex) in
- let mkOp T retT op_code args := constr:(Op (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) (t1:=T) (tR:=retT) op_code args) in
- let mkMatchPair tC ex eC := constr:(MatchPair (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) (tC:=tC) ex eC) in
- let mkFst ex := constr:(Fst (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) ex) in
- let mkSnd ex := constr:(Snd (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) ex) in
- let reify_pretag := constr:(@exprf base_type_code interp_base_type op) in
- let reify_tag := constr:(reify_pretag var) in
- let dummy := debug_enter_reifyf e in
- match goal with
- | [ re := ?rev : reify reify_tag e |- _ ] =>
- (* fast path *)
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case hyp) in
- let ret := rev in
- let dummy := debug_leave_reifyf_success e ret in
- ret
- | _ =>
- let ret :=
- lazymatch e with
- | let x := ?ex in @?eC x =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case let_in) in
- let ex := reify_rec ex in
- let eC := reify_rec eC in
- mkLetIn ex eC
- | (dlet x := ?ex in @?eC x) =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case dlet_in) in
- let ex := reify_rec ex in
- let eC := reify_rec eC in
- mkLetIn ex eC
- | pair ?a ?b =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case pair) in
- let a := reify_rec a in
- let b := reify_rec b in
- mkPair a b
- | (fun x : ?T => ?C) =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case function) in
- let t := reify_flat_type T in
- (* Work around Coq 8.5 and 8.6 bug *)
- (* <https://coq.inria.fr/bugs/show_bug.cgi?id=4998> *)
- (* Avoid re-binding the Gallina variable referenced by Ltac [x] *)
- (* even if its Gallina name matches a Ltac in this tactic. *)
- let not_x := refresh x in
- let C' := match constr:(Set) with
- | _ => constr:(fun (x : T) (not_x : var t) (_ : reify_var_for_in_is base_type_code x t not_x) =>
- (_ : reify_internal reify_tag C)) (* [C] here is an open term that references "x" by name *)
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Error: reifyf: Failed to reify by typeclasses:" e)
- end in
- match constr:(Set) with
- | _ => lazymatch C'
- with fun _ v _ => @?C v => C end
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Error: reifyf: Failed to eliminate function dependencies of:" C')
- end
- | match ?ev with pair a b => @?eC a b end =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case matchpair) in
- let T := type of eC in
- let t := (let T := match (eval cbv beta in T) with _ -> _ -> ?T => T end in reify_flat_type T) in
- let v := reify_rec ev in
- let C := reify_rec eC in
- let ret := mkMatchPair t v C in
- ret
- | @fst ?A ?B ?ev =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case fst) in
- let v := reify_rec ev in
- mkFst v
- | @snd ?A ?B ?ev =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case snd) in
- let v := reify_rec ev in
- mkSnd v
- | ?x =>
- let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case generic) in
- let t := type of x in
- let t := reify_flat_type t in
- let retv := match constr:(Set) with
- | _ => let retv := reifyf_var x mkVar in constr:(finished_value retv)
- | _ => let op_head := head x in
- reify_op op op_head x
- | _ => lazymatch x with
- | ?F ?args
- => lazymatch goal with
- | [ rF : forall x not_x, reify reify_tag (F x) |- _ ]
- => constr:(context_value rF args)
- | [ rF : forall var' x (not_x : var' _), reify (reify_pretag var') (F x) |- _ ]
- => constr:(context_value (rF var) args)
- end
- end
- | _ => let c := mkConst t x in
- constr:(finished_value c)
- | _ => constr:(reification_unsuccessful)
- end in
- lazymatch retv with
- | finished_value ?v => v
- | context_value ?rFH ?eargs
- => let dummy := debug_reifyf_case ltac:(fun _ => idtac_reifyf_case context_value) in
- let args := reify_rec eargs in
- let F_head := head rFH in
- let F := lazymatch (eval cbv beta delta [F_head] in rFH) with
- | fun _ => ?C => C
- end in
- mkLetIn args F
- | op_info (reify_op _ _ ?nargs ?op_code)
- =>
- let tR := (let tR := type of x in reify_flat_type tR) in
- lazymatch nargs with
- | 1%nat
- => lazymatch x with
- | ?f ?x0
- => let a0T := (let t := type of x0 in reify_flat_type t) in
- let a0 := reify_rec x0 in
- mkOp a0T tR op_code a0
- end
- | 2%nat
- => lazymatch x with
- | ?f ?x0 ?x1
- => let a0T := (let t := type of x0 in reify_flat_type t) in
- let a0 := reify_rec x0 in
- let a1T := (let t := type of x1 in reify_flat_type t) in
- let a1 := reify_rec x1 in
- let args := mkPair a0 a1 in
- mkOp (@Prod _ a0T a1T) tR op_code args
- end
- | 3%nat
- => lazymatch x with
- | ?f ?x0 ?x1 ?x2
- => let a0T := (let t := type of x0 in reify_flat_type t) in
- let a0 := reify_rec x0 in
- let a1T := (let t := type of x1 in reify_flat_type t) in
- let a1 := reify_rec x1 in
- let a2T := (let t := type of x2 in reify_flat_type t) in
- let a2 := reify_rec x2 in
- let args := let a01 := mkPair a0 a1 in mkPair a01 a2 in
- mkOp (@Prod _ (@Prod _ a0T a1T) a2T) tR op_code args
- end
- | 4%nat
- => lazymatch x with
- | ?f ?x0 ?x1 ?x2 ?x3
- => let a0T := (let t := type of x0 in reify_flat_type t) in
- let a0 := reify_rec x0 in
- let a1T := (let t := type of x1 in reify_flat_type t) in
- let a1 := reify_rec x1 in
- let a2T := (let t := type of x2 in reify_flat_type t) in
- let a2 := reify_rec x2 in
- let a3T := (let t := type of x3 in reify_flat_type t) in
- let a3 := reify_rec x3 in
- let args := let a01 := mkPair a0 a1 in let a012 := mkPair a01 a2 in mkPair a012 a3 in
- mkOp (@Prod _ (@Prod _ (@Prod _ a0T a1T) a2T) a3T) tR op_code args
- end
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Error: Unsupported number of operation arguments in reifyf:" nargs)
- end
- | reification_unsuccessful
- => constr_run_tac_fail ltac:(fun _ => idtac "Error: Failed to reify:" x)
- end
- end in
- let dummy := debug_leave_reifyf_success e ret in
- ret
- | _ => debug_leave_reifyf_failure e
- end.
-
-Ltac do_reifyf_goal _ :=
- debug_enter_reify_rec;
- let e := lazymatch goal with
- | [ |- reify (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e ]
- => reifyf base_type_code interp_base_type op var e
- | [ |- reify_internal (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e ]
- => reifyf base_type_code interp_base_type op var e
- end in
- debug_leave_reify_rec e;
- eexact e.
-
-Hint Extern 0 (reify (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e)
-=> do_reifyf_goal () : typeclass_instances.
-Hint Extern 0 (reify_internal (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e)
-=> do_reifyf_goal () : typeclass_instances.
-
-(** For reification including [Abs] *)
-Class reify_abs_internal {varT} (var : varT) {eT} (e : eT) {T : Type} := Build_reify_abs_internal : T.
-Class reify_abs {varT} (var : varT) {eT} (e : eT) {T : Type} := Build_reify_abs : T.
-Typeclasses Opaque reify_abs_internal reify_abs.
-Ltac reify_abs base_type_code interp_base_type op var e :=
- let reify_rec e := reify_abs base_type_code interp_base_type op var e in
- let reifyf_term e := reifyf base_type_code interp_base_type op var e in
- let mkReturn ef := constr:(Return (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) ef) in
- let mkAbs src ef := constr:(Abs (base_type_code:=base_type_code) (interp_base_type:=interp_base_type) (op:=op) (var:=var) (src:=src) ef) in
- let reify_pretag := constr:(@exprf base_type_code interp_base_type op) in
- let reify_tag := constr:(reify_pretag var) in
- let dummy := debug_enter_reify_abs e in
- lazymatch goal with
- | [ re := ?rev : forall var' (x : ?T) (not_x : var' _), reify (reify_pretag var') (e x) |- _ ] =>
- (* fast path *)
- let t := reify_flat_type T in
- let F := lazymatch (eval cbv beta in (rev var)) with
- | fun _ => ?C => C
- end in
- mkAbs t F
- | _ =>
- lazymatch e with
- | (fun x : ?T => ?C) =>
- let t := reify_flat_type T in
- let not_x := refresh x in
- let C' := match constr:(Set) with
- | _ => constr:(fun (x : T) (not_x : var t) (_ : reify_var_for_in_is base_type_code x t not_x) =>
- (_ : reify_abs_internal reify_tag C)) (* [C] here is an open term that references "x" by name *)
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Error: reify_abs: Failed to reify by typeclasses:" e)
- end in
- let C := match constr:(Set) with
- | _ => lazymatch C'
- with fun _ v _ => @?C v => C end
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Error: reify_abs: Failed to eliminate function dependencies of:" C')
- end in
- mkAbs t C
- | ?x =>
- let xv := reifyf_term x in
- mkReturn xv
- end
- end.
-
-Ltac do_reify_abs_goal _ :=
- debug_enter_reify_rec;
- let e := lazymatch goal with
- | [ |- reify_abs (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e ]
- => reify_abs base_type_code interp_base_type op var e
- | [ |- reify_abs_internal (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e ]
- => reify_abs base_type_code interp_base_type op var e
- end in
- debug_leave_reify_rec e;
- eexact e.
-
-Hint Extern 0 (reify_abs (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e)
-=> do_reify_abs_goal () : typeclass_instances.
-Hint Extern 0 (reify_abs_internal (@exprf ?base_type_code ?interp_base_type ?op ?var) ?e)
-=> do_reify_abs_goal () : typeclass_instances.
-
-Ltac Reify' base_type_code interp_base_type op e :=
- lazymatch constr:(fun (var : flat_type base_type_code -> Type) => (_ : reify_abs_internal (@exprf base_type_code interp_base_type op var) e)) with
- (fun var => ?C) => constr:(fun (var : flat_type base_type_code -> Type) => C) (* copy the term but not the type cast *)
- end.
-Ltac Reify base_type_code interp_base_type op make_const e :=
- let r := Reify' base_type_code interp_base_type op e in
- let r := lazymatch type of r with
- | forall var, exprf _ _ _ _ => constr:(fun var => Abs (src:=Unit) (fun _ => r var))
- | _ => r
- end in
- constr:(@InputSyntax.Compile base_type_code interp_base_type op make_const _ r).
-
-Ltac rhs_of_goal :=
- lazymatch goal with
- | [ |- ?R ?LHS ?RHS ] => RHS
- | [ |- forall x, ?R (@?LHS x) (@?RHS x) ] => RHS
- end.
-
-Ltac transitivity_tt term :=
- first [ transitivity term
- | transitivity (term tt)
- | let x := fresh in intro x; transitivity (term x); revert x ].
-
-Ltac Reify_rhs_gen Reify prove_interp_compile_correct interp_op try_tac :=
- Tuple.unfold_flat_interp_tuple ();
- let rhs := rhs_of_goal in
- let RHS := Reify rhs in
- let RHS' := (eval vm_compute in RHS) in
- transitivity_tt (Syntax.Interp interp_op RHS');
- [
- | transitivity_tt (Syntax.Interp interp_op RHS);
- [ lazymatch goal with
- | [ |- ?R ?x ?y ]
- => cut (x = y)
- | [ |- forall k, ?R (?x k) (?y k) ]
- => cut (x = y)
- end;
- [ let H := fresh in
- intro H; rewrite H; reflexivity
- | apply f_equal; vm_compute; reflexivity ]
- | intros; etransitivity; (* first we strip off the [InputSyntax.Compile]
- bit; Coq is bad at inferring the type, so we
- help it out by providing it *)
- [ cbv [InputSyntax.compilet];
- prove_interp_compile_correct ()
- | try_tac
- ltac:(fun _
- => (* now we unfold the interpretation function,
- including the parameterized bits; we assume that
- [hnf] is enough to unfold the interpretation
- functions that we're parameterized over. *)
- clear;
- abstract (
- lazymatch goal with
- | [ |- context[@InputSyntax.Interp ?base_type_code ?interp_base_type ?op ?interp_op ?t ?e] ]
- => let interp_base_type' := (eval hnf in interp_base_type) in
- let interp_op' := (eval hnf in interp_op) in
- change interp_base_type with interp_base_type';
- change interp_op with interp_op'
- end;
- subst_let;
- cbv iota beta delta [InputSyntax.Interp interp_type interp_type_gen interp_type_gen_hetero interp_flat_type interp interpf InputSyntax.Fst InputSyntax.Snd]; reflexivity)) ] ] ].
-
-Ltac prove_compile_correct_using tac :=
- fun _ => intros;
- lazymatch goal with
- | [ |- @Syntax.Interp ?base_type_code ?interp_base_type ?op ?interp_op _ (@Compile _ _ _ ?make_const (InputSyntax.Arrow ?src (Tflat ?dst)) ?e) ?x = _ ]
- => apply (fun pf => @InputSyntax.Compile_correct base_type_code interp_base_type op make_const interp_op pf src dst e x);
- solve [ tac () ]
- | [ |- @Syntax.Interp ?base_type_code ?interp_base_type ?op ?interp_op _ (@Compile _ _ _ ?make_const (Tflat ?T) ?e) ?x = _ ]
- => apply (fun pf => @InputSyntax.Compile_flat_correct_flat base_type_code interp_base_type op make_const interp_op pf T e x);
- solve [ tac () ]
- end.
-Ltac prove_compile_correct :=
- prove_compile_correct_using
- ltac:(fun _ => let T := fresh in intro T; destruct T; reflexivity).
-
-Ltac Reify_rhs base_type_code interp_base_type op make_const interp_op :=
- Reify_rhs_gen
- ltac:(Reify base_type_code interp_base_type op make_const)
- prove_compile_correct
- interp_op
- ltac:(fun tac => tac ()).
-
-Definition Reify_evar_package
- {base_type_code interp_base_type op}
- (make_const : forall t : base_type_code,
- interp_base_type t -> op ()%ctype (Tbase t))
- (interp_op : forall s d : flat_type base_type_code, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d)
- {t}
- (f : Syntax.interp_type interp_base_type t)
- := evar_rel_package
- f (Compilers.Syntax.Expr _ op t)
- (fun e f
- => forall x, Compilers.Syntax.Interp interp_op e x = f x).
-
-Definition Interp_Reify_evar_package
- {base_type_code interp_base_type op make_const interp_op t f}
- (pkg : @Reify_evar_package base_type_code interp_base_type op make_const interp_op t f)
- : forall x, Compilers.Syntax.Interp interp_op (val pkg) x = f x
- := evar_package_pf pkg.
-
-Ltac autosolve else_tac :=
- lazymatch goal with
- | [ |- @Reify_evar_package ?base_type_code ?interp_base_type ?op ?make_const ?interp_op _ _ ]
- => eexists; cbv beta; Reify_rhs base_type_code interp_base_type op make_const interp_op; reflexivity
- | _ => else_tac ()
- end.
-
-Ltac SideConditions.CorePackages.autosolve ::= autosolve.
-
-(** Reification of context variables of the form [F := _ :
- Syntax.interp_type _ _] *)
-Ltac unique_reify_context_variable base_type_code interp_base_type op F Fbody rT :=
- let reify_pretag := constr:(@exprf base_type_code interp_base_type op) in
- lazymatch goal with
- | [ H : forall var x not_x, reify _ (F x) |- _ ]
- => fail
- | _
- => let H' := fresh in
- let src := lazymatch rT with Syntax.Arrow ?src ?dst => src end in
- lazymatch Fbody with
- | fun x : ?X => ?Fbody'
- => let not_x := refresh x in
- let rF := lazymatch constr:(fun var' (x : X) (not_x : var' src) (_ : reify_var_for_in_is base_type_code x src not_x)
- => (_ : reify_internal (reify_pretag var') Fbody'))
- with
- | fun (var' : ?VAR) (x : ?X) (v : ?V) _ => ?C
- => constr:(fun (var' : VAR) (v : V) => C)
- end in
- let F' := fresh F in
- pose rF as F';
- pose ((fun var (x : X) => F' var) : forall var (x : X) (not_x : var src), reify (reify_pretag var) (F x)) as H';
- cbv beta in (value of H')
- end
- end.
-Ltac prereify_context_variables interp_base_type :=
- (** N.B. this assumes that [interp_base_type] is a transparent
- definition; minor reorganization may be needed if this is changed
- (moving the burden of reifying [interp_base_type T] to
- [reify_base_type], rather than keeping it here) *)
- cbv beta iota delta [interp_base_type] in *.
-Ltac reify_context_variable base_type_code interp_base_type op :=
- (** [match reverse] so that we respect the chain of dependencies in
- context variables; otherwise we're going to be trying the last
- context variable many times, and bottlenecking there. *)
- match reverse goal with
- | [ F := ?Fbody : Syntax.interp_type _ ?rT |- _ ]
- => unique_reify_context_variable base_type_code interp_base_type op F Fbody rT
- end.
-Ltac lazy_reify_context_variable base_type_code interp_base_type op :=
- lazymatch reverse goal with
- | [ F := ?Fbody : Syntax.interp_type _ ?rT |- _ ]
- => unique_reify_context_variable base_type_code interp_base_type op F Fbody rT
- end.
-Ltac reify_context_variables base_type_code interp_base_type op :=
- prereify_context_variables interp_base_type;
- repeat reify_context_variable base_type_code interp_base_type op.
diff --git a/src/Compilers/Relations.v b/src/Compilers/Relations.v
deleted file mode 100644
index 27a101e4f..000000000
--- a/src/Compilers/Relations.v
+++ /dev/null
@@ -1,368 +0,0 @@
-Require Import Coq.Lists.List Coq.Classes.RelationClasses Coq.Classes.Morphisms.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Sigma.
-
-Local Coercion is_true : bool >-> Sortclass.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
-
- Local Ltac rel_relb_t :=
- repeat first [ progress simpl in *
- | reflexivity
- | intuition congruence
- | setoid_rewrite Bool.andb_true_iff
- | intro
- | rewrite_hyp <- !* ].
-
- Section flat_type.
- Context {interp_base_type1 interp_base_type2 : base_type_code -> Type}.
- Local Notation interp_flat_type1 := (interp_flat_type interp_base_type1).
- Local Notation interp_flat_type2 := (interp_flat_type interp_base_type2).
-
- Section gen_Prop.
- Context (P : Type)
- (and : P -> P -> P)
- (True : P)
- (False : P).
- Section pointwise1.
- Context (R : forall t, interp_base_type1 t -> P).
- Fixpoint interp_flat_type_rel_pointwise1_gen_Prop (t : flat_type)
- : interp_flat_type1 t -> P :=
- match t with
- | Tbase t => R t
- | Unit => fun _ => True
- | Prod A B => fun x : interp_flat_type _ A * interp_flat_type _ B
- => and (interp_flat_type_rel_pointwise1_gen_Prop _ (fst x))
- (interp_flat_type_rel_pointwise1_gen_Prop _ (snd x))
- end.
- End pointwise1.
- Section pointwise2.
- Context (R : forall t, interp_base_type1 t -> interp_base_type2 t -> P).
- Fixpoint interp_flat_type_rel_pointwise_gen_Prop (t : flat_type)
- : interp_flat_type1 t -> interp_flat_type2 t -> P :=
- match t with
- | Tbase t => R t
- | Unit => fun _ _ => True
- | Prod A B
- => fun (x : interp_flat_type _ A * interp_flat_type _ B)
- (y : interp_flat_type _ A * interp_flat_type _ B)
- => and (interp_flat_type_rel_pointwise_gen_Prop _ (fst x) (fst y))
- (interp_flat_type_rel_pointwise_gen_Prop _ (snd x) (snd y))
- end.
- End pointwise2.
- Section pointwise2_hetero.
- Context (R : forall t1 t2, interp_base_type1 t1 -> interp_base_type2 t2 -> P).
- Fixpoint interp_flat_type_rel_pointwise_hetero_gen_Prop (t1 t2 : flat_type)
- : interp_flat_type1 t1 -> interp_flat_type2 t2 -> P
- := match t1, t2 with
- | Tbase t1, Tbase t2 => R t1 t2
- | Unit, Unit => fun _ _ => True
- | Prod x1 y1, Prod x2 y2
- => fun (a b : interp_flat_type _ _ * interp_flat_type _ _)
- => and (interp_flat_type_rel_pointwise_hetero_gen_Prop x1 x2 (fst a) (fst b))
- (interp_flat_type_rel_pointwise_hetero_gen_Prop y1 y2 (snd a) (snd b))
- | Tbase _, _
- | Unit, _
- | Prod _ _, _
- => fun _ _ => False
- end.
- End pointwise2_hetero.
- End gen_Prop.
-
- Definition interp_flat_type_relb_pointwise1
- := @interp_flat_type_rel_pointwise1_gen_Prop bool andb true.
- Global Arguments interp_flat_type_relb_pointwise1 _ !_ _ / .
- Definition interp_flat_type_rel_pointwise1
- := @interp_flat_type_rel_pointwise1_gen_Prop Prop and True.
- Global Arguments interp_flat_type_rel_pointwise1 _ !_ _ / .
- Lemma interp_flat_type_rel_pointwise1_iff_relb {R} t x
- : interp_flat_type_relb_pointwise1 R t x <-> interp_flat_type_rel_pointwise1 R t x.
- Proof using Type. clear; induction t; rel_relb_t. Qed.
- Definition interp_flat_type_rel_pointwise1_gen_Prop_iff_bool
- : forall {R} t x,
- interp_flat_type_rel_pointwise1_gen_Prop bool _ _ R t x
- <-> interp_flat_type_rel_pointwise1_gen_Prop Prop _ _ R t x
- := @interp_flat_type_rel_pointwise1_iff_relb.
- Definition interp_flat_type_relb_pointwise
- := @interp_flat_type_rel_pointwise_gen_Prop bool andb true.
- Global Arguments interp_flat_type_relb_pointwise _ !_ _ _ / .
- Definition interp_flat_type_rel_pointwise
- := @interp_flat_type_rel_pointwise_gen_Prop Prop and True.
- Global Arguments interp_flat_type_rel_pointwise _ !_ _ _ / .
- Lemma interp_flat_type_rel_pointwise_iff_relb {R} t x y
- : interp_flat_type_relb_pointwise R t x y <-> interp_flat_type_rel_pointwise R t x y.
- Proof using Type. clear; induction t; rel_relb_t. Qed.
- Definition interp_flat_type_rel_pointwise_gen_Prop_iff_bool
- : forall {R} t x y,
- interp_flat_type_rel_pointwise_gen_Prop bool _ _ R t x y
- <-> interp_flat_type_rel_pointwise_gen_Prop Prop _ _ R t x y
- := @interp_flat_type_rel_pointwise_iff_relb.
- Definition interp_flat_type_relb_pointwise_hetero
- := @interp_flat_type_rel_pointwise_hetero_gen_Prop bool andb true false.
- Global Arguments interp_flat_type_relb_pointwise_hetero _ !_ !_ _ _ / .
- Definition interp_flat_type_rel_pointwise_hetero
- := @interp_flat_type_rel_pointwise_hetero_gen_Prop Prop and True False.
- Global Arguments interp_flat_type_rel_pointwise_hetero _ !_ !_ _ _ / .
- Lemma interp_flat_type_rel_pointwise_hetero_iff_relb {R} t1 t2 x y
- : interp_flat_type_relb_pointwise_hetero R t1 t2 x y <-> interp_flat_type_rel_pointwise_hetero R t1 t2 x y.
- Proof using Type. clear; revert dependent t2; induction t1, t2; rel_relb_t. Qed.
- Definition interp_flat_type_rel_pointwise_hetero_gen_Prop_iff_bool
- : forall {R} t1 t2 x y,
- interp_flat_type_rel_pointwise_hetero_gen_Prop bool _ _ _ R t1 t2 x y
- <-> interp_flat_type_rel_pointwise_hetero_gen_Prop Prop _ _ _ R t1 t2 x y
- := @interp_flat_type_rel_pointwise_hetero_iff_relb.
-
- Lemma interp_flat_type_rel_pointwise_hetero_iff {R t} x y
- : interp_flat_type_rel_pointwise (fun t => R t t) t x y
- <-> interp_flat_type_rel_pointwise_hetero R t t x y.
- Proof using Type. induction t; simpl; rewrite_hyp ?*; reflexivity. Qed.
-
- Lemma interp_flat_type_rel_pointwise_impl {R1 R2 : forall t, _ -> _ -> Prop} t x y
- : interp_flat_type_rel_pointwise (fun t x y => (R1 t x y -> R2 t x y)%type) t x y
- -> (interp_flat_type_rel_pointwise R1 t x y
- -> interp_flat_type_rel_pointwise R2 t x y).
- Proof using Type. induction t; simpl; intuition. Qed.
-
- Lemma interp_flat_type_rel_pointwise_always {R : forall t, _ -> _ -> Prop}
- : (forall t x y, R t x y)
- -> forall t x y, interp_flat_type_rel_pointwise R t x y.
- Proof using Type. induction t; simpl; intuition. Qed.
- End flat_type.
- Section flat_type_extra.
- Context {interp_base_type1 interp_base_type2 : base_type_code -> Type}.
- Lemma interp_flat_type_rel_pointwise_impl' {R1 R2 : forall t, _ -> _ -> Prop} t x y
- : @interp_flat_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => (R1 t y x -> R2 t x y)%type) t x y
- -> (interp_flat_type_rel_pointwise R1 t y x
- -> interp_flat_type_rel_pointwise R2 t x y).
- Proof using Type. induction t; simpl; intuition. Qed.
-
- Global Instance interp_flat_type_rel_pointwise_Reflexive {R : forall t, _ -> _ -> Prop} {H : forall t, Reflexive (R t)}
- : forall t, Reflexive (@interp_flat_type_rel_pointwise interp_base_type1 interp_base_type1 R t).
- Proof using Type.
- induction t; intro; simpl; try apply conj; try reflexivity.
- Qed.
-
- Lemma interp_flat_type_rel_pointwise_SmartVarfMap
- {interp_base_type1' interp_base_type2'}
- {R : forall t, _ -> _ -> Prop}
- t f g x y
- : @interp_flat_type_rel_pointwise interp_base_type1 interp_base_type2 R t (SmartVarfMap f x) (SmartVarfMap g y)
- <-> @interp_flat_type_rel_pointwise interp_base_type1' interp_base_type2' (fun t x y => R t (f _ x) (g _ y)) t x y.
- Proof using Type.
- induction t; simpl; try reflexivity.
- rewrite_hyp <- !*; reflexivity.
- Qed.
- End flat_type_extra.
-
- Section type.
- Section hetero.
- Context (interp_src1 interp_src2 : flat_type -> Type)
- (interp_dst1 interp_dst2 : flat_type -> Type).
- Section hetero.
- Context (Rsrc : forall t, interp_src1 t -> interp_src2 t -> Prop)
- (Rdst : forall t, interp_dst1 t -> interp_dst2 t -> Prop).
-
- Definition interp_type_gen_rel_pointwise_hetero (t : type)
- : interp_type_gen_hetero interp_src1 interp_dst1 t
- -> interp_type_gen_hetero interp_src2 interp_dst2 t
- -> Prop
- := @respectful_hetero _ _ _ _ (Rsrc _) (fun _ _ => Rdst _).
- Global Arguments interp_type_gen_rel_pointwise_hetero _ _ _ / .
- End hetero.
- Section hetero_hetero.
- Context (Rsrc : forall t1 t2, interp_src1 t1 -> interp_src2 t2 -> Prop)
- (Rdst : forall t1 t2, interp_dst1 t1 -> interp_dst2 t2 -> Prop).
-
- Fixpoint interp_type_gen_rel_pointwise_hetero_hetero (t1 t2 : type)
- : interp_type_gen_hetero interp_src1 interp_dst1 t1
- -> interp_type_gen_hetero interp_src2 interp_dst2 t2
- -> Prop
- := @respectful_hetero _ _ _ _ (Rsrc _ _) (fun _ _ => Rdst _ _).
- Global Arguments interp_type_gen_rel_pointwise_hetero_hetero _ _ _ _ / .
- End hetero_hetero.
- End hetero.
-
- Section partially_hetero.
- Context (interp_flat_type1 interp_flat_type2 : flat_type -> Type)
- (R : forall t, interp_flat_type1 t -> interp_flat_type2 t -> Prop).
-
- Definition interp_type_gen_rel_pointwise
- : forall t,
- interp_type_gen interp_flat_type1 t
- -> interp_type_gen interp_flat_type2 t
- -> Prop
- := interp_type_gen_rel_pointwise_hetero
- interp_flat_type1 interp_flat_type2
- interp_flat_type1 interp_flat_type2
- R R.
- Global Arguments interp_type_gen_rel_pointwise _ _ _ / .
- End partially_hetero.
- End type.
-
- Section specialized_type.
- Section hetero.
- Context (interp_base_type1 interp_base_type2 : base_type_code -> Type).
- Definition interp_type_rel_pointwise R
- : forall t, interp_type interp_base_type1 t
- -> interp_type interp_base_type2 t
- -> Prop
- := interp_type_gen_rel_pointwise _ _ (interp_flat_type_rel_pointwise R).
- Global Arguments interp_type_rel_pointwise _ !_ _ _ / .
-
- Definition interp_type_rel_pointwise_hetero R
- : forall t1 t2, interp_type interp_base_type1 t1
- -> interp_type interp_base_type2 t2
- -> Prop
- := interp_type_gen_rel_pointwise_hetero_hetero _ _ _ _ (interp_flat_type_rel_pointwise_hetero R) (interp_flat_type_rel_pointwise_hetero R).
- Global Arguments interp_type_rel_pointwise_hetero _ !_ !_ _ _ / .
- End hetero.
- End specialized_type.
-
- Section lifting.
- Context {interp_base_type1 interp_base_type2 : base_type_code -> Type}.
- Local Notation interp_flat_type1 := (interp_flat_type interp_base_type1).
- Local Notation interp_flat_type2 := (interp_flat_type interp_base_type2).
- Let Tbase := (@Tbase base_type_code).
- Local Coercion Tbase : base_type_code >-> flat_type.
-
- Section with_rel.
- Context (R : forall t, interp_flat_type1 t -> interp_flat_type2 t -> Prop)
- (RUnit : R Unit tt tt).
- Section RProd.
- Context (RProd : forall A B x y, R A (fst x) (fst y) /\ R B (snd x) (snd y) -> R (Prod A B) x y)
- (RProd' : forall A B x y, R (Prod A B) x y -> R A (fst x) (fst y) /\ R B (snd x) (snd y)).
- Lemma lift_interp_flat_type_rel_pointwise1 t (x : interp_flat_type1 t) (y : interp_flat_type2 t)
- : interp_flat_type_rel_pointwise R t x y -> R t x y.
- Proof using RProd RUnit. clear RProd'; induction t; simpl; destruct_head_hnf' unit; intuition. Qed.
- Lemma lift_interp_flat_type_rel_pointwise2 t (x : interp_flat_type1 t) (y : interp_flat_type2 t)
- : R t x y -> interp_flat_type_rel_pointwise R t x y.
- Proof using RProd'. clear RProd; induction t; simpl; destruct_head_hnf' unit; split_and; intuition. Qed.
- End RProd.
- Section RProd_iff.
- Context (RProd : forall A B x y, R A (fst x) (fst y) /\ R B (snd x) (snd y) <-> R (Prod A B) x y).
- Lemma lift_interp_flat_type_rel_pointwise t (x : interp_flat_type1 t) (y : interp_flat_type2 t)
- : interp_flat_type_rel_pointwise R t x y <-> R t x y.
- Proof using RProd RUnit.
- split_iff; split; auto using lift_interp_flat_type_rel_pointwise1, lift_interp_flat_type_rel_pointwise2.
- Qed.
- End RProd_iff.
- End with_rel.
- Lemma lift_interp_flat_type_rel_pointwise_f_eq {T} (f g : forall t, _ -> T t) t x y
- : @interp_flat_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => f t x = g t y)
- t x y
- <-> SmartVarfMap f x = SmartVarfMap g y.
- Proof using Type.
- induction t; unfold SmartVarfMap in *; simpl in *; destruct_head_hnf unit; try tauto.
- rewrite_hyp !*; intuition congruence.
- Qed.
- Lemma lift_interp_flat_type_rel_pointwise_f_eq_id1 (f : forall t, _ -> _) t x y
- : @interp_flat_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => x = f t y)
- t x y
- <-> x = SmartVarfMap f y.
- Proof using Type. rewrite lift_interp_flat_type_rel_pointwise_f_eq, SmartVarfMap_id; reflexivity. Qed.
- Lemma lift_interp_flat_type_rel_pointwise_f_eq_id2 (f : forall t, _ -> _) t x y
- : @interp_flat_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => f t x = y)
- t x y
- <-> SmartVarfMap f x = y.
- Proof using Type. rewrite lift_interp_flat_type_rel_pointwise_f_eq, SmartVarfMap_id; reflexivity. Qed.
- Lemma lift_interp_flat_type_rel_pointwise_f_eq2 {T} (f g : forall t, _ -> _ -> T t) t x y
- : @interp_flat_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => f t x y = g t x y)
- t x y
- <-> SmartVarfMap2 f x y = SmartVarfMap2 g x y.
- Proof using Type.
- induction t; unfold SmartVarfMap2 in *; simpl in *; destruct_head_hnf unit; try tauto.
- rewrite_hyp !*; intuition congruence.
- Qed.
- Lemma lift_interp_type_rel_pointwise_f_eq {T} (f g : forall t, _ -> T t) t x y
- : interp_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => f t x = g t y)
- t x y
- <-> (forall a b, SmartVarfMap f a = SmartVarfMap g b -> SmartVarfMap f (x a) = SmartVarfMap g (y b)).
- Proof using Type.
- destruct t; simpl; unfold interp_type_rel_pointwise, respectful_hetero.
- setoid_rewrite lift_interp_flat_type_rel_pointwise_f_eq; reflexivity.
- Qed.
- Lemma lift_interp_type_rel_pointwise_f_eq_id1 (f : forall t, _ -> _) t x y
- : interp_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => x = f t y)
- t x y
- <-> (forall a, x (SmartVarfMap f a) = SmartVarfMap f (y a)).
- Proof using Type. rewrite lift_interp_type_rel_pointwise_f_eq; setoid_rewrite SmartVarfMap_id; firstorder (subst; eauto). Qed.
- Lemma lift_interp_type_rel_pointwise_f_eq_id2 (f : forall t, _ -> _) t x y
- : interp_type_rel_pointwise
- interp_base_type1 interp_base_type2
- (fun t x y => f t x = y)
- t x y
- <-> (forall a, SmartVarfMap f (x a) = y (SmartVarfMap f a)).
- Proof using Type. rewrite lift_interp_type_rel_pointwise_f_eq; setoid_rewrite SmartVarfMap_id; firstorder (subst; eauto). Qed.
- End lifting.
-
- Local Ltac t :=
- repeat match goal with
- | _ => intro
- | [ H : False |- _ ] => exfalso; assumption
- | _ => progress subst
- | _ => assumption
- | _ => progress inversion_sigma
- | _ => progress inversion_prod
- | _ => progress simpl in *
- | _ => progress destruct_head_hnf' and
- | [ H : context[List.In _ (_ ++ _)] |- _ ]
- => rewrite List.in_app_iff in H
- | _ => progress destruct_head' or
- | _ => solve [ eauto ]
- end.
-
- Lemma interp_flat_type_rel_pointwise_flatten_binding_list
- {interp_base_type1 interp_base_type2 t T} R' e1 e2 v1 v2
- (H : List.In (existT _ t (v1, v2)%core) (flatten_binding_list e1 e2))
- (HR : @interp_flat_type_rel_pointwise interp_base_type1 interp_base_type2 R' T e1 e2)
- : R' t v1 v2.
- Proof using Type. induction T; t. Qed.
-
- Lemma interp_flat_type_rel_pointwise_hetero_flatten_binding_list2
- {interp_base_type1 interp_base_type2 t1 t2 T1 T2} R' e1 e2 v1 v2
- (H : List.In (existT _ (t1, t2)%core (v1, v2)%core) (flatten_binding_list2 e1 e2))
- (HR : @interp_flat_type_rel_pointwise_hetero interp_base_type1 interp_base_type2 R' T1 T2 e1 e2)
- : R' t1 t2 v1 v2.
- Proof using Type.
- revert dependent T2; induction T1, T2; t.
- Qed.
-End language.
-
-Global Arguments interp_type_rel_pointwise {_ _ _} R {t} _ _.
-Global Arguments interp_type_rel_pointwise_hetero {_ _ _} R {t1 t2} _ _.
-Global Arguments interp_type_gen_rel_pointwise_hetero_hetero {_ _ _ _ _} Rsrc Rdst {t1 t2} _ _.
-Global Arguments interp_type_gen_rel_pointwise_hetero {_ _ _ _ _} Rsrc Rdst {t} _ _.
-Global Arguments interp_type_gen_rel_pointwise {_ _ _} R {t} _ _.
-Global Arguments interp_flat_type_rel_pointwise_gen_Prop {_ _ _ P} and True R {t} _ _.
-Global Arguments interp_flat_type_rel_pointwise_hetero_gen_Prop {_ _ _ P} and True False R {t1 t2} _ _.
-Global Arguments interp_flat_type_rel_pointwise_hetero {_ _ _} R {t1 t2} _ _.
-Global Arguments interp_flat_type_relb_pointwise_hetero {_ _ _} R {t1 t2} _ _.
-Global Arguments interp_flat_type_rel_pointwise1 {_ _} R {t} _.
-Global Arguments interp_flat_type_relb_pointwise1 {_ _} R {t} _.
-Global Arguments interp_flat_type_rel_pointwise {_ _ _} R {t} _ _.
-Global Arguments interp_flat_type_relb_pointwise {_ _ _} R {t} _ _.
diff --git a/src/Compilers/RenameBinders.v b/src/Compilers/RenameBinders.v
deleted file mode 100644
index 7be603cbe..000000000
--- a/src/Compilers/RenameBinders.v
+++ /dev/null
@@ -1,78 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-
-Ltac uncurry_f f :=
- let t := type of f in
- lazymatch eval compute in t with
- | prod ?a ?b -> ?R
- => uncurry_f (fun x y => f (@pair a b x y))
- | ?a -> ?R
- => let x := fresh in
- constr:(fun x : a => ltac:(let v := uncurry_f (f x) in exact v))
- | _ => f
- end.
-Ltac make_destruct_specialize t with_destruct_specialize_tac :=
- let do_tac T1 T2 n1 mk_n2 :=
- pose tt as n1;
- make_destruct_specialize
- T1
- ltac:(fun destruct_specialize_ab
- => let n2 := mk_n2 () in
- pose tt as n2;
- make_destruct_specialize
- T2
- ltac:(fun destruct_specialize_cd
- => with_destruct_specialize_tac
- ltac:(fun arg f cont =>
- clear n1 n2;
- refine (let '(n1, n2)%core := arg in _);
- clear arg;
- destruct_specialize_ab
- n1 f
- ltac:(fun f => destruct_specialize_cd n2 f cont)))) in
- lazymatch eval compute in t with
- | prod (prod ?a ?b) (prod ?c ?d)
- => let arg1 := fresh "arg" in
- do_tac (prod a b) (prod c d) arg1 ltac:(fun _ => fresh "arg")
- | prod (prod ?a ?b) ?c
- => let arg1 := fresh "arg" in
- do_tac (prod a b) c arg1 ltac:(fun _ => fresh "x")
- | prod ?a (prod ?c ?d)
- => let arg1 := fresh "x" in
- do_tac a (prod c d) arg1 ltac:(fun _ => fresh "arg")
- | prod ?a ?b
- => let arg1 := fresh "x" in
- do_tac a b arg1 ltac:(fun _ => fresh "x")
- | _
- => with_destruct_specialize_tac ltac:(fun arg f cont => cont (f arg))
- end.
-Ltac renamify input :=
- let t := type of input in
- let t := (eval compute in t) in
- let ret :=
- constr:(ltac:(
- let var := fresh "var" in
- intro var;
- let input := constr:(input var) in
- let input := (eval compute in input) in
- let arg := fresh "arg" in
- refine (Abs (fun arg => _));
- let input := constr:(invert_Abs input) in
- let t := type of arg in
- let t := (eval compute in t) in
- let input := uncurry_f input in
- let input := (eval cbv iota beta delta [invert_Abs] in input) in
- make_destruct_specialize
- t ltac:(fun do_destruct_specialize
- => do_destruct_specialize
- arg input
- ltac:(fun input => let input := (eval cbv beta in input) in
- exact input))
- ) : t) in
- (eval cbv beta zeta in ret).
-Notation renamify f :=
- (let t := _ in
- let renamify_F0 : t := f in
- ((fun renamify_F : t => ltac:(let v := renamify renamify_F in exact v))
- renamify_F0))
- (only parsing).
diff --git a/src/Compilers/Rewriter.v b/src/Compilers/Rewriter.v
deleted file mode 100644
index a23627480..000000000
--- a/src/Compilers/Rewriter.v
+++ /dev/null
@@ -1,39 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var : base_type_code -> Type}.
- Context (rewrite_op_expr : forall src dst (opc : op src dst),
- exprf (var:=var) src -> exprf (var:=var) dst).
-
- Fixpoint rewrite_opf {t} (e : @exprf var t) : @exprf var t
- := match e in Syntax.exprf _ _ t return @exprf var t with
- | LetIn tx ex tC eC
- => LetIn (@rewrite_opf tx ex) (fun x => @rewrite_opf tC (eC x))
- | Var _ x => Var x
- | TT => TT
- | Pair tx ex ty ey
- => Pair (@rewrite_opf tx ex) (@rewrite_opf ty ey)
- | Op t1 tR opc args => rewrite_op_expr _ _ opc (@rewrite_opf t1 args)
- end.
-
- Definition rewrite_op {t} (e : @expr var t) : @expr var t
- := match e in Syntax.expr _ _ t return @expr var t with
- | Abs _ _ f => Abs (fun x => rewrite_opf (f x))
- end.
- End with_var.
-
- Definition RewriteOp
- (rewrite_op_expr : forall var src dst, op src dst -> @exprf var src -> @exprf var dst)
- {t} (e : Expr t)
- : Expr t
- := fun var => rewrite_op (rewrite_op_expr _) (e _).
-End language.
diff --git a/src/Compilers/RewriterInterp.v b/src/Compilers/RewriterInterp.v
deleted file mode 100644
index 1f9fa5bed..000000000
--- a/src/Compilers/RewriterInterp.v
+++ /dev/null
@@ -1,50 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Rewriter.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-
-Section language.
- Context {base_type_code : Type}
- {interp_base_type : base_type_code -> Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Local Notation exprf := (@exprf base_type_code op interp_base_type).
- Local Notation expr := (@expr base_type_code op interp_base_type).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section specialized.
- Context {rewrite_op_expr : forall src dst (opc : op src dst), exprf src -> exprf dst}
- (Hrewrite : forall src dst opc args,
- interpf interp_op (rewrite_op_expr src dst opc args)
- = interp_op _ _ opc (interpf interp_op args)).
-
- Lemma interpf_rewrite_opf {t} (e : exprf t)
- : interpf interp_op (rewrite_opf rewrite_op_expr e) = interpf interp_op e.
- Proof using Type*.
- induction e; simpl; unfold LetIn.Let_In; rewrite_hyp ?*; reflexivity.
- Qed.
-
- Lemma interp_rewrite_op {t} (e : expr t)
- : forall x, interp interp_op (rewrite_op rewrite_op_expr e) x = interp interp_op e x.
- Proof using Type*.
- destruct e; intro x; apply interpf_rewrite_opf.
- Qed.
- End specialized.
-
- Lemma InterpRewriteOp
- {rewrite_op_expr}
- (Hrewrite : forall src dst opc args,
- interpf interp_op (rewrite_op_expr interp_base_type src dst opc args)
- = interp_op _ _ opc (interpf interp_op args))
- {t} (e : Expr t)
- : forall x, Interp interp_op (RewriteOp rewrite_op_expr e) x = Interp interp_op e x.
- Proof using Type.
- apply interp_rewrite_op; assumption.
- Qed.
-End language.
-
-Hint Rewrite @InterpRewriteOp using assumption : reflective_interp.
diff --git a/src/Compilers/RewriterWf.v b/src/Compilers/RewriterWf.v
deleted file mode 100644
index 22c135d00..000000000
--- a/src/Compilers/RewriterWf.v
+++ /dev/null
@@ -1,61 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Compilers.Rewriter.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
- Local Notation Wf := (@Wf base_type_code op).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}
- {rewrite_op_expr1 : forall src dst (opc : op src dst),
- exprf (var:=var1) src -> exprf (var:=var1) dst}
- {rewrite_op_expr2 : forall src dst (opc : op src dst),
- exprf (var:=var2) src -> exprf (var:=var2) dst}
- (Hrewrite_wf : forall G src dst opc args1 args2,
- wff G args1 args2
- -> wff G
- (rewrite_op_expr1 src dst opc args1)
- (rewrite_op_expr2 src dst opc args2)).
-
- Lemma wff_rewrite_opf {t} G (e1 : @exprf var1 t) (e2 : @exprf var2 t)
- (Hwf : wff G e1 e2)
- : wff G (rewrite_opf rewrite_op_expr1 e1) (rewrite_opf rewrite_op_expr2 e2).
- Proof using Type*.
- induction Hwf; simpl; try constructor; auto.
- Qed.
-
- Lemma wf_rewrite_op {t} (e1 : @expr var1 t) (e2 : @expr var2 t)
- (Hwf : wf e1 e2)
- : wf (rewrite_op rewrite_op_expr1 e1) (rewrite_op rewrite_op_expr2 e2).
- Proof using Type*.
- destruct Hwf; simpl; constructor; intros; apply wff_rewrite_opf; auto.
- Qed.
- End with_var.
-
- Lemma Wf_RewriteOp
- {rewrite_op_expr}
- (Hrewrite_wff : forall var1 var2 G src dst opc args1 args2,
- wff G args1 args2
- -> wff G
- (rewrite_op_expr var1 src dst opc args1)
- (rewrite_op_expr var2 src dst opc args2))
- {t} (e : Expr t)
- (Hwf : Wf e)
- : Wf (RewriteOp rewrite_op_expr e).
- Proof using Type.
- intros var1 var2; apply wf_rewrite_op; auto.
- Qed.
-End language.
-
-Hint Resolve Wf_RewriteOp : wf.
diff --git a/src/Compilers/SmartMap.v b/src/Compilers/SmartMap.v
deleted file mode 100644
index b02bba2a1..000000000
--- a/src/Compilers/SmartMap.v
+++ /dev/null
@@ -1,453 +0,0 @@
-Require Import Coq.Classes.Morphisms.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Tuple.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Notations.
-
-Section homogenous_type.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {var : base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation interp_flat_type := (@interp_flat_type base_type_code).
- Local Notation exprf := (@exprf base_type_code op var).
- Local Notation expr := (@expr base_type_code op var).
-
- (** Sometimes, we want to deal with partially-interpreted
- expressions, things like [prod (exprf A) (exprf B)] rather than
- [exprf (Prod A B)], or like [prod (var A) (var B)] when we start
- with the type [Prod A B]. These convenience functions let us
- recurse on the type in only one place, and replace one kind of
- pairing operator (be it [pair] or [Pair] or anything else) with
- another kind, and simultaneously mapping a function over the
- base values (e.g., [Var] (for turning [var] into [exprf]) or
- [Const] (for turning [interp_base_type] into [exprf])). *)
- Fixpoint smart_interp_flat_map {f g}
- (h : forall x, f x -> g (Tbase x))
- (tt : g Unit)
- (pair : forall A B, g A -> g B -> g (Prod A B))
- {t}
- : interp_flat_type f t -> g t
- := match t return interp_flat_type f t -> g t with
- | Syntax.Tbase _ => h _
- | Unit => fun _ => tt
- | Prod A B => fun v : interp_flat_type _ A * interp_flat_type _ B
- => pair _ _
- (@smart_interp_flat_map f g h tt pair A (fst v))
- (@smart_interp_flat_map f g h tt pair B (snd v))
- end.
- Fixpoint smart_interp_flat_map2 {f1 f2 g}
- (h : forall x, f1 x -> f2 x -> g (Tbase x))
- (tt : g Unit)
- (pair : forall A B, g A -> g B -> g (Prod A B))
- {t}
- : interp_flat_type f1 t -> interp_flat_type f2 t -> g t
- := match t return interp_flat_type f1 t -> interp_flat_type f2 t -> g t with
- | Syntax.Tbase _ => h _
- | Unit => fun _ _ => tt
- | Prod A B => fun (v1 : interp_flat_type _ A * interp_flat_type _ B)
- (v2 : interp_flat_type _ A * interp_flat_type _ B)
- => pair _ _
- (@smart_interp_flat_map2 f1 f2 g h tt pair A (fst v1) (fst v2))
- (@smart_interp_flat_map2 f1 f2 g h tt pair B (snd v1) (snd v2))
- end.
- Fixpoint smart_interp_flat_map3 {f1 f2 f3 g}
- (h : forall x, f1 x -> f2 x -> f3 x -> g (Tbase x))
- (tt : g Unit)
- (pair : forall A B, g A -> g B -> g (Prod A B))
- {t}
- : interp_flat_type f1 t -> interp_flat_type f2 t -> interp_flat_type f3 t -> g t
- := match t return interp_flat_type f1 t -> interp_flat_type f2 t -> interp_flat_type f3 t -> g t with
- | Syntax.Tbase _ => h _
- | Unit => fun _ _ _ => tt
- | Prod A B => fun (v1 : interp_flat_type _ A * interp_flat_type _ B)
- (v2 : interp_flat_type _ A * interp_flat_type _ B)
- (v3 : interp_flat_type _ A * interp_flat_type _ B)
- => pair _ _
- (@smart_interp_flat_map3 f1 f2 f3 g h tt pair A (fst v1) (fst v2) (fst v3))
- (@smart_interp_flat_map3 f1 f2 f3 g h tt pair B (snd v1) (snd v2) (snd v3))
- end.
- Definition smart_interp_map_hetero {f g g'}
- (h : forall x, f x -> g (Tbase x))
- (tt : g Unit)
- (pair : forall A B, g A -> g B -> g (Prod A B))
- (abs : forall A B, (g A -> g B) -> g' (Arrow A B))
- {t}
- : interp_type_gen_hetero g (interp_flat_type f) t -> g' t
- := match t return interp_type_gen_hetero g (interp_flat_type f) t -> g' t with
- | Arrow A B => fun v => abs _ _
- (fun x => @smart_interp_flat_map f g h tt pair _ (v x))
- end.
- Fixpoint SmartValf {T} (val : forall t : base_type_code, T t) t : interp_flat_type T t
- := match t return interp_flat_type T t with
- | Syntax.Tbase _ => val _
- | Unit => tt
- | Prod A B => (@SmartValf T val A, @SmartValf T val B)
- end.
- Section SmartValf_monad.
- Context (M : Type -> Type) (ret : forall T, T -> M T)
- (bind : forall A B, M A -> (A -> M B) -> M B).
- Fixpoint SmartValfM
- {T} (val : forall t : base_type_code, M (T t)) t : M (interp_flat_type T t)
- := match t return M (interp_flat_type T t) with
- | Syntax.Tbase _ => val _
- | Unit => ret _ tt
- | Prod A B => bind _ _ (@SmartValfM T val A)
- (fun a => bind _ _ (@SmartValfM T val B)
- (fun b => ret _ (a, b)))
- end.
- End SmartValf_monad.
-
- (** [SmartVar] is like [Var], except that it inserts
- pair-projections and [Pair] as necessary to handle [flat_type],
- and not just [base_type_code] *)
- Local Notation exprfb := (fun t => exprf (Tbase t)).
- Definition SmartValf_option {T} (val : forall t, option (T t)) t
- : option (interp_flat_type T t)
- := @SmartValfM
- (fun t => option t) (fun t v => @Some t v)
- (fun _ _ x f => match x with
- | Some x => f x
- | None => None
- end)
- T val t.
- Definition SmartPairf {t} : interp_flat_type exprfb t -> exprf t
- := @smart_interp_flat_map exprfb exprf (fun t x => x) TT (fun A B x y => Pair x y) t.
- Lemma SmartPairf_Pair {A B} (e1 : interp_flat_type _ A) (e2 : interp_flat_type _ B)
- : SmartPairf (t:=Prod A B) (e1, e2)%core = Pair (SmartPairf e1) (SmartPairf e2).
- Proof using Type. reflexivity. Qed.
- Definition SmartVarf {t} : interp_flat_type var t -> exprf t
- := @smart_interp_flat_map var exprf (fun t => Var) TT (fun A B x y => Pair x y) t.
- Definition SmartVarf_Pair {A B v}
- : @SmartVarf (Prod A B) v = Pair (SmartVarf (fst v)) (SmartVarf (snd v))
- := eq_refl.
- Definition SmartVarfMap {var var'} (f : forall t, var t -> var' t) {t}
- : interp_flat_type var t -> interp_flat_type var' t
- := @smart_interp_flat_map var (interp_flat_type var') f tt (fun A B x y => pair x y) t.
- Lemma SmartVarfMap_compose {var' var'' var''' t} f g x
- : @SmartVarfMap var'' var''' g t (@SmartVarfMap var' var'' f t x)
- = @SmartVarfMap _ _ (fun t v => g t (f t v)) t x.
- Proof using Type.
- unfold SmartVarfMap; clear; induction t; simpl; destruct_head_hnf unit; destruct_head_hnf prod;
- rewrite_hyp ?*; congruence.
- Qed.
- Lemma SmartVarfMap_id {var' t} x : @SmartVarfMap var' var' (fun _ x => x) t x = x.
- Proof using Type.
- unfold SmartVarfMap; clear; induction t; simpl; destruct_head_hnf unit; destruct_head_hnf prod;
- rewrite_hyp ?*; congruence.
- Qed.
- Definition SmartVarfMap_Pair {var' var''} {f' : forall t, var' t -> var'' t} {A B}
- v
- : @SmartVarfMap var' var'' f' (Prod A B) v
- = (SmartVarfMap f' (fst v), SmartVarfMap f' (snd v))
- := eq_refl.
- Lemma SmartVarfMap_tuple' {var' var''} {f' : forall t, var' t -> var'' t} {T n}
- v
- : @SmartVarfMap var' var'' f' (tuple' T n) v
- = flat_interp_untuple' (Tuple.map' (@SmartVarfMap var' var'' f' _) (flat_interp_tuple' v)).
- Proof.
- induction n as [|n IHn]; [ reflexivity | destruct v as [v0 v1] ].
- simpl; rewrite SmartVarfMap_Pair, IHn; simpl.
- reflexivity.
- Qed.
- Definition SmartVarfMap_tuple {var' var''} {f' : forall t, var' t -> var'' t} {T n}
- v
- : @SmartVarfMap var' var'' f' (tuple T n) v
- = tuple_map (@SmartVarfMap var' var'' f' _) v.
- Proof.
- destruct n as [|n]; [ destruct v; reflexivity | ].
- apply SmartVarfMap_tuple'.
- Qed.
- Global Instance smart_interp_flat_map_Proper {f g}
- : Proper ((forall_relation (fun t => pointwise_relation _ eq))
- ==> eq
- ==> (forall_relation (fun A => forall_relation (fun B => pointwise_relation _ (pointwise_relation _ eq))))
- ==> forall_relation (fun t => eq ==> eq))
- (@smart_interp_flat_map f g).
- Proof using Type.
- unfold forall_relation, pointwise_relation, respectful.
- intros F G HFG x y ? Q R HQR t a b ?; subst y b.
- induction t; simpl in *; auto.
- rewrite_hyp !*; reflexivity.
- Qed.
- Global Instance SmartVarfMap_Proper {var' var''}
- : Proper (forall_relation (fun t => pointwise_relation _ eq) ==> forall_relation (fun t => eq ==> eq))
- (@SmartVarfMap var' var'').
- Proof using Type.
- repeat intro; eapply smart_interp_flat_map_Proper; trivial; repeat intro; reflexivity.
- Qed.
- Definition SmartVarfMap2 {var var' var''} (f : forall t, var t -> var' t -> var'' t) {t}
- : interp_flat_type var t -> interp_flat_type var' t -> interp_flat_type var'' t
- := @smart_interp_flat_map2 var var' (interp_flat_type var'') f tt (fun A B x y => pair x y) t.
- Lemma SmartVarfMap2_fst_arg {var' var''} {t}
- (x : interp_flat_type var' t)
- (y : interp_flat_type var'' t)
- : SmartVarfMap2 (fun _ a b => a) x y = x.
- Proof using Type.
- unfold SmartVarfMap2; clear; induction t; simpl; destruct_head_hnf unit; destruct_head_hnf prod;
- rewrite_hyp ?*; congruence.
- Qed.
- Lemma SmartVarfMap2_snd_arg {var' var''} {t}
- (x : interp_flat_type var' t)
- (y : interp_flat_type var'' t)
- : SmartVarfMap2 (fun _ a b => b) x y = y.
- Proof using Type.
- unfold SmartVarfMap2; clear; induction t; simpl; destruct_head_hnf unit; destruct_head_hnf prod;
- rewrite_hyp ?*; congruence.
- Qed.
- Definition SmartVarfMap3 {var var' var'' var'''} (f : forall t, var t -> var' t -> var'' t -> var''' t) {t}
- : interp_flat_type var t -> interp_flat_type var' t -> interp_flat_type var'' t -> interp_flat_type var''' t
- := @smart_interp_flat_map3 var var' var'' (interp_flat_type var''') f tt (fun A B x y => pair x y) t.
- Definition SmartVarfTypeMap {var} (f : forall t, var t -> Type) {t}
- : interp_flat_type var t -> Type
- := @smart_interp_flat_map var (fun _ => Type) f unit (fun _ _ P Q => P * Q)%type t.
- Definition SmartVarfPropMap {var} (f : forall t, var t -> Prop) {t}
- : interp_flat_type var t -> Prop
- := @smart_interp_flat_map var (fun _ => Prop) f True (fun _ _ P Q => P /\ Q)%type t.
- Definition SmartVarfTypeMap2 {var var'} (f : forall t, var t -> var' t -> Type) {t}
- : interp_flat_type var t -> interp_flat_type var' t -> Type
- := @smart_interp_flat_map2 var var' (fun _ => Type) f unit (fun _ _ P Q => P * Q)%type t.
- Definition SmartVarfPropMap2 {var var'} (f : forall t, var t -> var' t -> Prop) {t}
- : interp_flat_type var t -> interp_flat_type var' t -> Prop
- := @smart_interp_flat_map2 var var' (fun _ => Prop) f True (fun _ _ P Q => P /\ Q)%type t.
- Definition SmartFlatTypeMap {var'} (f : forall t, var' t -> base_type_code) {t}
- : interp_flat_type var' t -> flat_type
- := @smart_interp_flat_map var' (fun _ => flat_type) (fun t v => Tbase (f t v)) Unit (fun _ _ => Prod) t.
- Definition SmartFlatTypeMap_Pair {var'} (f : forall t, var' t -> base_type_code) {A B}
- (x : interp_flat_type var' (A * B))
- : SmartFlatTypeMap f x
- = (SmartFlatTypeMap f (@fst (interp_flat_type _ _) (interp_flat_type _ _) x)
- * SmartFlatTypeMap f (@snd (interp_flat_type _ _) (interp_flat_type _ _) x))%ctype
- := eq_refl.
- Definition SmartFlatTypeUnMap (t : flat_type)
- : interp_flat_type (fun _ => base_type_code) t
- := SmartValf (fun t => t) t.
- Fixpoint SmartFlatTypeMapInterp {var' var''} (f : forall t, var' t -> base_type_code)
- (fv : forall t v, var'' (f t v)) t {struct t}
- : forall v, interp_flat_type var'' (SmartFlatTypeMap f (t:=t) v)
- := match t return forall v, interp_flat_type var'' (SmartFlatTypeMap f (t:=t) v) with
- | Syntax.Tbase x => fv _
- | Unit => fun v => v
- | Prod A B => fun xy : interp_flat_type _ A * interp_flat_type _ B
- => (@SmartFlatTypeMapInterp _ _ f fv A (fst xy),
- @SmartFlatTypeMapInterp _ _ f fv B (snd xy))
- end.
- Fixpoint SmartFlatTypeMapInterp2 {var' var'' var'''} (f : forall t, var' t -> base_type_code)
- (fv : forall t v, var'' t -> var''' (f t v)) t {struct t}
- : forall v, interp_flat_type var'' t -> interp_flat_type var''' (SmartFlatTypeMap f (t:=t) v)
- := match t return forall v, interp_flat_type var'' t -> interp_flat_type var''' (SmartFlatTypeMap f (t:=t) v) with
- | Syntax.Tbase x => fv _
- | Unit => fun v _ => v
- | Prod A B => fun (xy : interp_flat_type _ A * interp_flat_type _ B)
- (x'y' : interp_flat_type _ A * interp_flat_type _ B)
- => (@SmartFlatTypeMapInterp2 _ _ _ f fv A (fst xy) (fst x'y'),
- @SmartFlatTypeMapInterp2 _ _ _ f fv B (snd xy) (snd x'y'))
- end.
- Fixpoint SmartFlatTypeMapUnInterp var' var'' var''' (f : forall t, var' t -> base_type_code)
- (fv : forall t (v : var' t), var'' (f t v) -> var''' t)
- {t} {struct t}
- : forall v, interp_flat_type var'' (SmartFlatTypeMap f (t:=t) v)
- -> interp_flat_type var''' t
- := match t return forall v, interp_flat_type var'' (SmartFlatTypeMap f (t:=t) v)
- -> interp_flat_type var''' t with
- | Syntax.Tbase x => fv _
- | Unit => fun _ v => v
- | Prod A B => fun (v : interp_flat_type _ A * interp_flat_type _ B)
- (xy : interp_flat_type _ (SmartFlatTypeMap _ (fst v)) * interp_flat_type _ (SmartFlatTypeMap _ (snd v)))
- => (@SmartFlatTypeMapUnInterp _ _ _ f fv A _ (fst xy),
- @SmartFlatTypeMapUnInterp _ _ _ f fv B _ (snd xy))
- end.
- Definition SmartVarMap {var' var''} (f : forall t, var' t -> var'' t) (f' : forall t, var'' t -> var' t) {t}
- : interp_type_gen (interp_flat_type var') t -> interp_type_gen (interp_flat_type var'') t
- := match t return interp_type_gen (interp_flat_type var') t -> interp_type_gen (interp_flat_type var'') t with
- | Arrow src dst => fun F x => SmartVarfMap f (F (SmartVarfMap f' x))
- end.
- Lemma SmartVarMap_id {var' t} x v : @SmartVarMap var' var' (fun _ x => x) (fun _ x => x) t x v = x v.
- Proof using Type. destruct t; simpl; rewrite !SmartVarfMap_id; reflexivity. Qed.
- Definition SmartVarVarf {t} : interp_flat_type var t -> interp_flat_type exprfb t
- := SmartVarfMap (fun t => Var).
- Definition SmartVarVarf_Pair {A B} (v : interp_flat_type _ _ * interp_flat_type _ _)
- : @SmartVarVarf (Prod A B) v
- = (SmartVarVarf (fst v), SmartVarVarf (snd v))
- := eq_refl.
- Lemma SmartPairfSmartVarVarf_SmartVarf {t} v
- : SmartPairf (SmartVarVarf v) = SmartVarf (t:=t) v.
- Proof.
- induction t; try reflexivity; simpl.
- rewrite SmartVarf_Pair, SmartVarVarf_Pair, SmartPairf_Pair; f_equal;
- auto.
- Qed.
-End homogenous_type.
-
-Global Arguments SmartVarf {_ _ _ _} _.
-Global Arguments SmartPairf {_ _ _ t} _.
-Global Arguments SmartValf {_} T _ t.
-Global Arguments SmartVarVarf {_ _ _ _} _.
-Global Arguments SmartVarfMap {_ _ _} _ {!_} _ / .
-Global Arguments SmartVarfMap2 {_ _ _ _} _ {!t} _ _ / .
-Global Arguments SmartVarfMap3 {_ _ _ _ _} _ {!t} _ _ _ / .
-Global Arguments SmartVarfTypeMap {_ _} _ {_} _.
-Global Arguments SmartVarfPropMap {_ _} _ {_} _.
-Global Arguments SmartVarfTypeMap2 {_ _ _} _ {t} _ _.
-Global Arguments SmartVarfPropMap2 {_ _ _} _ {t} _ _.
-Global Arguments SmartFlatTypeMap {_ _} _ {_} _.
-Global Arguments SmartFlatTypeUnMap {_} _.
-Global Arguments SmartFlatTypeMapInterp {_ _ _ _} _ {_} _.
-Global Arguments SmartFlatTypeMapInterp2 {_ _ _ _ f} fv {t} _ _.
-Global Arguments SmartFlatTypeMapUnInterp {_ _ _ _ _} fv {_ _} _.
-Global Arguments SmartVarMap {_ _ _} _ _ {!_} _ / _.
-
-Section hetero_type.
- Fixpoint flatten_flat_type {base_type_code} (t : flat_type (flat_type base_type_code)) : flat_type base_type_code
- := match t with
- | Tbase T => T
- | Unit => Unit
- | Prod A B => Prod (@flatten_flat_type _ A) (@flatten_flat_type _ B)
- end.
-
- Section smart_flat_type_map2.
- Context {base_type_code1 base_type_code2 : Type}.
-
- Definition SmartFlatTypeMap2 {var' : base_type_code1 -> Type} (f : forall t, var' t -> flat_type base_type_code2) {t}
- : interp_flat_type var' t -> flat_type base_type_code2
- := @smart_interp_flat_map base_type_code1 var' (fun _ => flat_type base_type_code2) f Unit (fun _ _ => Prod) t.
- Fixpoint SmartFlatTypeMap2Interp {var' var''} (f : forall t, var' t -> flat_type base_type_code2)
- (fv : forall t v, interp_flat_type var'' (f t v)) t {struct t}
- : forall v, interp_flat_type var'' (SmartFlatTypeMap2 f (t:=t) v)
- := match t return forall v, interp_flat_type var'' (SmartFlatTypeMap2 f (t:=t) v) with
- | Tbase x => fv _
- | Unit => fun v => v
- | Prod A B => fun xy : interp_flat_type _ A * interp_flat_type _ B
- => (@SmartFlatTypeMap2Interp _ _ f fv A (fst xy),
- @SmartFlatTypeMap2Interp _ _ f fv B (snd xy))
- end.
- Fixpoint SmartFlatTypeMapUnInterp2 var' var'' var''' (f : forall t, var' t -> flat_type base_type_code2)
- (fv : forall t (v : var' t), interp_flat_type var'' (f t v) -> var''' t)
- {t} {struct t}
- : forall v, interp_flat_type var'' (SmartFlatTypeMap2 f (t:=t) v)
- -> interp_flat_type var''' t
- := match t return forall v, interp_flat_type var'' (SmartFlatTypeMap2 f (t:=t) v)
- -> interp_flat_type var''' t with
- | Tbase x => fv _
- | Unit => fun _ v => v
- | Prod A B => fun (v : interp_flat_type _ A * interp_flat_type _ B)
- (xy : interp_flat_type _ (SmartFlatTypeMap2 _ (fst v)) * interp_flat_type _ (SmartFlatTypeMap2 _ (snd v)))
- => (@SmartFlatTypeMapUnInterp2 _ _ _ f fv A _ (fst xy),
- @SmartFlatTypeMapUnInterp2 _ _ _ f fv B _ (snd xy))
- end.
- Fixpoint SmartFlatTypeMap2Interp2 {var' var'' var'''} (f : forall t, var' t -> flat_type base_type_code2)
- (fv : forall t v, var'' t -> interp_flat_type var''' (f t v)) t {struct t}
- : forall v, interp_flat_type var'' t -> interp_flat_type var''' (SmartFlatTypeMap2 f (t:=t) v)
- := match t return forall v, interp_flat_type var'' t -> interp_flat_type var''' (SmartFlatTypeMap2 f (t:=t) v) with
- | Tbase x => fv _
- | Unit => fun v _ => v
- | Prod A B => fun (xy : interp_flat_type _ A * interp_flat_type _ B)
- (x'y' : interp_flat_type _ A * interp_flat_type _ B)
- => (@SmartFlatTypeMap2Interp2 _ _ _ f fv A (fst xy) (fst x'y'),
- @SmartFlatTypeMap2Interp2 _ _ _ f fv B (snd xy) (snd x'y'))
- end.
-
- Lemma SmartFlatTypeMapUnInterp2_SmartFlatTypeMap2Interp2
- var' var'' var'''
- (f : forall t, var' t -> flat_type base_type_code2)
- (fv : forall t (v : var' t), interp_flat_type var'' (f t v) -> var''' t)
- (gv : forall t v, var''' t -> interp_flat_type var'' (f t v))
- {t} v
- (e : interp_flat_type var''' t)
- : @SmartFlatTypeMapUnInterp2
- _ _ _ f fv t v
- (@SmartFlatTypeMap2Interp2
- _ _ _ f gv t v e)
- = SmartVarfMap2 (fun t v e => fv t v (gv t v e)) v e.
- Proof using Type.
- induction t; simpl in *; destruct_head' unit;
- rewrite_hyp ?*; reflexivity.
- Qed.
- End smart_flat_type_map2.
-
- Section smart_flat_type.
- Context {base_type_code1 base_type_code2 : Type}
- (f : base_type_code1 -> base_type_code2).
- Fixpoint lift_flat_type (t : flat_type base_type_code1)
- : flat_type base_type_code2
- := match t with
- | Tbase T => Tbase (f T)
- | Unit => Unit
- | Prod A B => Prod (lift_flat_type A) (lift_flat_type B)
- end.
-
- Section with_var.
- Context {var1 : base_type_code1 -> Type}
- {var2 : base_type_code2 -> Type}
- (fvar : forall t, var1 t -> var2 (f t))
- (fvar' : forall t, var2 (f t) -> var1 t).
-
- Fixpoint transfer_interp_flat_type {t}
- : interp_flat_type var1 t
- -> interp_flat_type var2 (lift_flat_type t)
- := match t with
- | Tbase T => fvar _
- | Unit => fun v => v
- | Prod A B => fun ab : interp_flat_type _ A * interp_flat_type _ B
- => (@transfer_interp_flat_type _ (fst ab),
- @transfer_interp_flat_type _ (snd ab))%core
- end.
-
- Fixpoint untransfer_interp_flat_type {t}
- : interp_flat_type var2 (lift_flat_type t)
- -> interp_flat_type var1 t
- := match t with
- | Tbase T => fvar' _
- | Unit => fun v => v
- | Prod A B => fun ab : interp_flat_type _ (lift_flat_type A)
- * interp_flat_type _ (lift_flat_type B)
- => (@untransfer_interp_flat_type _ (fst ab),
- @untransfer_interp_flat_type _ (snd ab))%core
- end.
- End with_var.
- End smart_flat_type.
-End hetero_type.
-
-Global Arguments SmartFlatTypeMap2 {_ _ _} _ {!_} _ / .
-Global Arguments SmartFlatTypeMap2Interp {_ _ _ _ _} fv {_} _.
-Global Arguments SmartFlatTypeMap2Interp2 {_ _ _ _ _ _} fv {t} v _.
-Global Arguments SmartFlatTypeMapUnInterp2 {_ _ _ _ _ _} fv {_ _} _.
-
-Section interp_lemmas.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}
- {interp_base_type : base_type_code -> Type}
- {interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}.
-
- Local Notation exprfb := (fun t => exprf _ op (Tbase t)).
-
- Lemma interpf_SmartVarf
- {t} (e : interp_flat_type _ t)
- : @interpf _ interp_base_type _ interp_op _ (SmartVarf (var:=interp_base_type) e)
- = e.
- Proof.
- induction t as [ t | | A IHA B IHB ]; try destruct e; try reflexivity.
- rewrite !SmartVarf_Pair; cbn; rewrite IHA, IHB.
- reflexivity.
- Qed.
-
- Lemma interpf_SmartPairf'
- {t} (e : interp_flat_type exprfb t)
- : @interpf _ interp_base_type _ interp_op _ (SmartPairf e)
- = SmartVarfMap (fun t => interpf interp_op) e.
- Proof.
- induction t as [ t | | A IHA B IHB ]; try reflexivity.
- { destruct e.
- rewrite !SmartPairf_Pair, !SmartVarfMap_Pair, <- !IHA, <- !IHB.
- reflexivity. }
- Qed.
-
- Lemma interpf_SmartPairf
- {t} (e : interp_flat_type exprfb t)
- : @interpf _ interp_base_type _ interp_op _ (SmartPairf (var:=interp_base_type) e)
- = SmartVarfMap (fun t => interpf interp_op) e.
- Proof. apply interpf_SmartPairf'. Qed.
-End interp_lemmas.
diff --git a/src/Compilers/StripExpr.v b/src/Compilers/StripExpr.v
deleted file mode 100644
index 57b3e2365..000000000
--- a/src/Compilers/StripExpr.v
+++ /dev/null
@@ -1,35 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation exprf := (@exprf base_type_code op).
-
- Section with_var.
- Context {var : base_type_code -> Type}.
-
- Fixpoint strip_exprf {T} (e : @exprf (fun t => @exprf var (Tbase t)) T)
- : @exprf var T
- := match e with
- | TT => TT
- | Var t v
- => v
- | Op src dst opv args
- => Op opv (strip_exprf args)
- | LetIn tx ex tC eC
- => LetIn (strip_exprf ex)
- (fun v => strip_exprf (eC (SmartVarVarf v)))
- | Pair tx ex ty ey
- => Pair (strip_exprf ex) (strip_exprf ey)
- end.
-
- Definition strip_expr {T} (e : @expr (fun t => @exprf var (Tbase t)) T)
- : @expr var T
- := match e with
- | Abs src dst f => Abs (fun x => strip_exprf (f (SmartVarVarf x)))
- end.
- End with_var.
-End language.
diff --git a/src/Compilers/Syntax.v b/src/Compilers/Syntax.v
deleted file mode 100644
index c86567c6d..000000000
--- a/src/Compilers/Syntax.v
+++ /dev/null
@@ -1,154 +0,0 @@
-(** * PHOAS Representation of Gallina *)
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Notations.
-
-(** We parameterize the language over a type of basic type codes (for
- things like [Z], [word], [bool]), as well as a type of n-ary
- operations returning one value, and n-ary operations returning two
- values. *)
-Delimit Scope ctype_scope with ctype.
-Local Open Scope ctype_scope.
-Delimit Scope expr_scope with expr.
-Local Open Scope expr_scope.
-Section language.
- Context (base_type_code : Type).
-
- Inductive flat_type := Tbase (T : base_type_code) | Unit | Prod (A B : flat_type).
- Bind Scope ctype_scope with flat_type.
-
- Inductive type := Arrow (A : flat_type) (B : flat_type).
- Bind Scope ctype_scope with type.
-
- Infix "*" := Prod : ctype_scope.
- Notation "A -> B" := (Arrow A B) : ctype_scope.
- Local Coercion Tbase : base_type_code >-> flat_type.
-
- Section tuple.
- Context (T : flat_type).
- Fixpoint tuple' n :=
- match n with
- | O => T
- | S n' => (tuple' n' * T)%ctype
- end.
- Definition tuple n :=
- match n with
- | O => Unit
- | S n' => tuple' n'
- end.
- End tuple.
-
- Definition domain (t : type) : flat_type
- := match t with Arrow src dst => src end.
- Definition codomain (t : type) : flat_type
- := match t with Arrow src dst => dst end.
-
- Section interp.
- Definition interp_type_gen_hetero (interp_src interp_dst : flat_type -> Type) (t : type) :=
- (interp_src match t with Arrow x y => x end -> interp_dst match t with Arrow x y => y end)%type.
- Definition interp_type_gen (interp_flat_type : flat_type -> Type)
- := interp_type_gen_hetero interp_flat_type interp_flat_type.
- Section flat_type.
- Context (interp_base_type : base_type_code -> Type).
- Fixpoint interp_flat_type (t : flat_type) :=
- match t with
- | Tbase t => interp_base_type t
- | Unit => unit
- | Prod x y => prod (interp_flat_type x) (interp_flat_type y)
- end.
- Definition interp_type := interp_type_gen interp_flat_type.
- End flat_type.
- End interp.
-
- Section expr_param.
- Context (interp_base_type : base_type_code -> Type).
- Context (op : flat_type (* input tuple *) -> flat_type (* output type *) -> Type).
- Local Notation interp_type := (interp_type interp_base_type).
- Local Notation interp_flat_type_gen := interp_flat_type.
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
- Section expr.
- Context {var : base_type_code -> Type}.
-
- (** N.B. [Let] binds the components of a pair to separate variables, and does so recursively *)
- Inductive exprf : flat_type -> Type :=
- | TT : exprf Unit
- | Var {t} (v : var t) : exprf t
- | Op {t1 tR} (opc : op t1 tR) (args : exprf t1) : exprf tR
- | LetIn {tx} (ex : exprf tx) {tC} (eC : interp_flat_type_gen var tx -> exprf tC) : exprf tC
- | Pair {tx} (ex : exprf tx) {ty} (ey : exprf ty) : exprf (Prod tx ty).
- Bind Scope expr_scope with exprf.
- Inductive expr : type -> Type :=
- | Abs {src dst} (f : interp_flat_type_gen var src -> exprf dst) : expr (Arrow src dst).
- Bind Scope expr_scope with expr.
- End expr.
-
- Definition Expr (t : type) := forall var, @expr var t.
-
- Section interp.
- Context (interp_op : forall src dst, op src dst -> interp_flat_type src -> interp_flat_type dst).
-
- Definition interpf_step
- (interpf : forall {t} (e : @exprf interp_flat_type t), interp_flat_type t)
- {t} (e : @exprf interp_flat_type t) : interp_flat_type t
- := match e in exprf t return interp_flat_type t with
- | TT => tt
- | Var _ x => x
- | Op _ _ op args => @interp_op _ _ op (@interpf _ args)
- | LetIn _ ex _ eC => dlet x := @interpf _ ex in @interpf _ (eC x)
- | Pair _ ex _ ey => (@interpf _ ex, @interpf _ ey)
- end.
-
- Fixpoint interpf {t} e
- := @interpf_step (@interpf) t e.
-
- Definition interp {t} (e : @expr interp_base_type t) : interp_type t
- := fun x
- => @interpf
- _
- (match e in @expr _ t
- return interp_flat_type (domain t)
- -> exprf (codomain t)
- with
- | Abs _ _ f => f
- end x).
-
- Definition Interp {t} (E : Expr t) : interp_type t := interp (E _).
- End interp.
- End expr_param.
-End language.
-Global Arguments tuple' {_}%type_scope _%ctype_scope _%nat_scope.
-Global Arguments tuple {_}%type_scope _%ctype_scope _%nat_scope.
-Global Arguments Unit {_}%type_scope.
-Global Arguments Prod {_}%type_scope (_ _)%ctype_scope.
-Global Arguments Arrow {_}%type_scope (_ _)%ctype_scope.
-Global Arguments Tbase {_}%type_scope _%ctype_scope.
-Global Arguments domain {_}%type_scope _%ctype_scope.
-Global Arguments codomain {_}%type_scope _%ctype_scope.
-
-Global Arguments Var {_ _ _ _} _.
-Global Arguments TT {_ _ _}.
-Global Arguments Op {_ _ _ _ _} _ _.
-Global Arguments LetIn {_ _ _ _} _ {_} _.
-Global Arguments Pair {_ _ _ _} _ {_} _.
-Global Arguments Abs {_ _ _ _ _} _.
-Global Arguments interp_type_gen_hetero {_} _ _ _.
-Global Arguments interp_type_gen {_} _ _.
-Global Arguments interp_flat_type {_} _ _.
-Global Arguments interp_type {_} _ _.
-Global Arguments Interp {_ _ _} interp_op {t} _ _.
-Global Arguments interp {_ _ _} interp_op {t} _ _.
-Global Arguments interpf {_ _ _} interp_op {t} _.
-Global Arguments interp _ _ _ _ _ !_ / _.
-
-Module Export Notations.
- Notation "()" := (@Unit _) : ctype_scope.
- Notation "A * B" := (@Prod _ A B) : ctype_scope.
- Notation "A -> B" := (@Arrow _ A B) : ctype_scope.
- Notation "'slet' x .. y := A 'in' b" := (LetIn A%expr (fun x => .. (fun y => b%expr) .. )) : expr_scope.
- Notation "'λ' x .. y , t" := (Abs (fun x => .. (Abs (fun y => t%expr)) ..)) : expr_scope.
- Notation "( x , y , .. , z )" := (Pair .. (Pair x%expr y%expr) .. z%expr) : expr_scope.
- Notation "( )" := TT : expr_scope.
- Notation "()" := TT : expr_scope.
- Infix "^" := (@tuple _) : ctype_scope.
- Bind Scope ctype_scope with flat_type.
- Bind Scope ctype_scope with type.
-End Notations.
diff --git a/src/Compilers/TestCase.v b/src/Compilers/TestCase.v
deleted file mode 100644
index 858c09a26..000000000
--- a/src/Compilers/TestCase.v
+++ /dev/null
@@ -1,276 +0,0 @@
-Require Import Coq.omega.Omega Coq.micromega.Psatz.
-Require Import Coq.PArith.BinPos Coq.Lists.List.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Compilers.Named.RegisterAssign.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Equality.
-Require Export Crypto.Compilers.Reify.
-Require Import Crypto.Compilers.InputSyntax.
-Require Import Crypto.Compilers.CommonSubexpressionElimination.
-Require Crypto.Compilers.Linearize Crypto.Compilers.Inline.
-Require Import Crypto.Compilers.WfReflective.
-Require Import Crypto.Compilers.Conversion.
-Require Import Crypto.Util.NatUtil.
-Require Import Crypto.Util.ZRange.
-
-Import ReifyDebugNotations.
-
-Local Set Boolean Equality Schemes.
-Local Set Decidable Equality Schemes.
-Inductive base_type := Tnat.
-Definition interp_base_type (v : base_type) : Type :=
- match v with
- | Tnat => nat
- end.
-Local Notation tnat := (Tbase Tnat).
-Inductive op : flat_type base_type -> flat_type base_type -> Type :=
-| Const (v : nat) : op Unit tnat
-| Add : op (Prod tnat tnat) tnat
-| Mul : op (Prod tnat tnat) tnat
-| Sub : op (Prod tnat tnat) tnat.
-Notation "' x" := (Syntax.Op (Const x) _) (only printing, at level 10) : expr_scope.
-Notation "' x" := (Syntax.Var x) (only printing, at level 10) : expr_scope.
-Notation "a + b" := (Syntax.Op Add (a, b)%expr) : expr_scope.
-Notation "a * b" := (Syntax.Op Mul (a, b)%expr) : expr_scope.
-Notation "a - b" := (Syntax.Op Sub (a, b)%expr) : expr_scope.
-Notation "' x" := (Named.Op (Const x) _) (only printing, at level 10) : nexpr_scope.
-Notation "a + b" := (Named.Op Add (a, b)%nexpr) : nexpr_scope.
-Notation "a * b" := (Named.Op Mul (a, b)%nexpr) : nexpr_scope.
-Notation "a - b" := (Named.Op Sub (a, b)%nexpr) : nexpr_scope.
-Definition is_const s d (v : op s d) : bool := match v with Const _ => true | _ => false 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 with
- | Const v => fun _ => v
- | Add => fun xy => fst xy + snd xy
- | Mul => fun xy => fst xy * snd xy
- | Sub => fun xy => fst xy - snd xy
- end%nat.
-
-Global Instance: reify_op op plus 2 Add := I.
-Global Instance: reify_op op mult 2 Mul := I.
-Global Instance: reify_op op minus 2 Sub := I.
-Global Instance: reify type nat := Tnat.
-
-Definition make_const (t : base_type) : interp_base_type t -> op Unit (Tbase t)
- := match t with
- | Tnat => fun v => Const v
- end.
-Ltac Reify' e := Reify.Reify' base_type interp_base_type op e.
-Ltac Reify e := Reify.Reify base_type interp_base_type op make_const e.
-Ltac Reify_rhs := Reify.Reify_rhs base_type interp_base_type op make_const interp_op.
-Ltac reify_context_variables :=
- Reify.reify_context_variables base_type interp_base_type op.
-
-(*Ltac reify_debug_level ::= constr:(2).*)
-
-Goal (flat_type base_type -> Type) -> False.
- intro var.
- let x := reifyf base_type interp_base_type op var 1%nat in pose x.
- let x := Reify' 1%nat in unify x (fun var => Return (InputSyntax.Const (interp_base_type:=interp_base_type) (var:=var) (t:=Tbase Tnat) (op:=op) 1)).
- let x := reifyf base_type interp_base_type op var (1 + 1)%nat in pose x.
- let x := Reify' (1 + 1)%nat in unify x (fun var => Return (Op Add (Pair (InputSyntax.Const (interp_base_type:=interp_base_type) (var:=var) (t:=Tbase Tnat) (op:=op) 1) (InputSyntax.Const (interp_base_type:=interp_base_type) (var:=var) (t:=Tbase Tnat) (op:=op) 1)))).
- let x := reify_abs base_type interp_base_type op var (fun x => x + 1)%nat in pose x.
- let x := Reify' (fun x => x + 1)%nat in unify x (fun var => Abs (fun y => Return (Op Add (Pair (Var y) (InputSyntax.Const (interp_base_type:=interp_base_type) (var:=var) (t:=Tbase Tnat) (op:=op) 1))))).
- let x := reifyf base_type interp_base_type op var (let '(a, b) := (1, 1) in a + b)%nat in pose x.
- let x := reifyf base_type interp_base_type op var (let '(a, b, c) := (1, 1, 1) in a + b + c)%nat in pose x.
- let x := Reify' (fun x => let '(a, b) := (1, 1) in a + x)%nat in let x := (eval vm_compute in x) in pose x.
- let x := Reify' (fun x => let '(a, b, c, (d, e), f) := x in a + b + c + d + e + f)%nat in let x := (eval vm_compute in x) in pose x.
- let x := Reify' (fun x => let '(a, b) := x in let '(a, c) := a in let '(a, d) := a in a + b + c + d)%nat in let x := (eval vm_compute in x) in pose x.
- let x := Reify' (fun ab0 : nat * nat * nat * nat => let (f, g6) := fst ab0 in
- let (f0, g7) := f in
- let ab3 := (1, 1) in
- let ab21 := (1, 1) in
- let z := snd ab3 + snd ab21 in z + z)%nat in let x := (eval vm_compute in x) in pose x.
- let x := Reify' (fun ab0 : nat * nat * nat => let (f, g7) := fst ab0 in 1 + 1) in let x := (eval vm_compute in x) in pose x.
- let x := Reify' (fun x => let '(a, b) := (1, 1) in a + x)%nat in
- unify x (fun var => Abs (fun x' =>
- let c1 := (InputSyntax.Const (interp_base_type:=interp_base_type) (var:=var) (t:=Tbase Tnat) (op:=op) 1) in
- Return (MatchPair (tC:=tnat) (Pair c1 c1)
- (fun x0 y0 : var tnat => Op Add (Pair (Var x0) (Var x')))))).
- let x := reifyf base_type interp_base_type op var (let x := 5 in let y := 6 in (let a := 1 in let '(c, d) := (2, 3) in a + x + c + d) + y)%nat in pose x.
- let x := Reify' (let x := 1 in let y := 1 in (let a := 1 in let '(c, d) := (2, 3) in a + x + c + d) + y)%nat in pose x.
- let x := Reify' (fun xy => let '(x, y) := xy in (let a := 1 in let '(c, d) := (2, 3) in a + x + c + d) + y)%nat in pose x.
-Abort.
-
-
-Goal (0 = let x := 1+2 in x*3)%nat.
- Reify_rhs.
-Abort.
-
-Goal (0 = let x := 1 in let y := 2 in x * y)%nat.
- Reify_rhs.
-Abort.
-
-Import Linearize Inline Eta.
-
-(** Example of flattening / linearization / linearize / a-normal form *)
-Goal True.
- let x := Reify (fun xy => let '(x, y) := xy in (let a := 1 in let '(c, d) := (2, 3) in let e := c + d in a + x - a + c + d + e) + y)%nat in
- pose x as e0;
- pose (ExprEta (Linearize x)) as e1;
- pose (ExprEta (InlineConst is_const (ANormal x))) as e.
- vm_compute in e0, e1, e.
-Abort.
-
-Definition example_expr : Syntax.Expr base_type op (Syntax.Arrow (tnat * tnat) tnat).
-Proof.
- let x := Reify (fun zw => let '(z, w) := zw in let unused := 1 + 1 in let x := 1 in let y := 1 in (let a := 1 in let cd := let cdef := (2, 3, 4, 5) in let '(c, d, e, f) := cdef in (c, d) in let '(c, d) := cd in a + x + (x + x) + (x + x) - (x + x) - a + c + d) + y + z + w)%nat in
- exact x.
-Defined.
-
-Definition example_expr_ctx : Syntax.Expr base_type op (Syntax.Arrow (tnat * tnat) tnat).
-Proof.
- pose (((fun ab => let '(a, b) := ab in a + b)%nat)
- : Syntax.interp_type interp_base_type (Syntax.Arrow (tnat * tnat) tnat))
- as F.
- reify_context_variables.
- let x := Reify (fun zw => let '(z, w) := zw in F (z, w))%nat in
- exact x.
-Defined.
-
-Definition base_type_eq_semidec_transparent : forall t1 t2, option (t1 = t2)
- := fun t1 t2 => match t1, t2 with
- | Tnat, Tnat => Some eq_refl
- end.
-Lemma base_type_eq_semidec_is_dec : forall t1 t2,
- base_type_eq_semidec_transparent t1 t2 = None -> t1 <> t2.
-Proof.
- intros t1 t2; destruct t1, t2; simpl; intros; congruence.
-Qed.
-Definition op_beq t1 tR : op t1 tR -> op t1 tR -> reified_Prop
- := fun x y => match x, y return bool with
- | Const a, Const b => NatUtil.nat_beq a b
- | Const _, _ => false
- | Add, Add => true
- | Add, _ => false
- | Mul, Mul => true
- | Mul, _ => false
- | Sub, Sub => true
- | Sub, _ => false
- end.
-Lemma op_beq_bl t1 tR (x y : op t1 tR)
- : to_prop (op_beq t1 tR x y) -> x = y.
-Proof.
- destruct x; simpl;
- refine match y with Add => _ | _ => _ end;
- repeat match goal with
- | _ => progress simpl in *
- | _ => progress unfold op_beq in *
- | [ |- context[reified_Prop_of_bool ?b] ]
- => destruct b eqn:?; unfold reified_Prop_of_bool
- | _ => progress nat_beq_to_eq
- | _ => congruence
- | _ => tauto
- end.
-Qed.
-
-Ltac reflect_Wf := WfReflective.reflect_Wf base_type_eq_semidec_is_dec op_beq_bl.
-
-Lemma example_expr_wf_slow : Wf example_expr.
-Proof.
- Time (vm_compute; intros;
- repeat match goal with
- | [ |- wf _ _ ] => constructor; intros
- | [ |- wff _ _ _ ] => constructor; intros
- | [ |- List.In _ _ ] => vm_compute
- | [ |- ?x = ?x \/ _ ] => left; reflexivity
- | [ |- ?x = ?y \/ _ ] => right
- end). (* 0.036 s *)
-Qed.
-
-Definition example_expr_eta := Eval vm_compute in Eta.ExprEta example_expr.
-
-Lemma example_expr_wf_eta : Wf example_expr_eta.
-Proof. Time reflect_Wf. (* 0.008 s *) Qed.
-
-Lemma example_expr_wf : Wf example_expr.
-Proof. Time reflect_Wf. (* 0.008 s *) Qed.
-
-Section cse.
- Let SConstT := nat.
- Inductive op_code : Set := SConst (v : nat) | SAdd | SMul | SSub.
- Definition symbolicify_op s d (v : op s d) : op_code
- := match v with
- | Const v => SConst v
- | Add => SAdd
- | Mul => SMul
- | Sub => SSub
- end.
- Definition CSE {t} e := @CSE base_type op_code base_type_beq op_code_beq internal_base_type_dec_bl op symbolicify_op (fun _ x => x) true t e (fun _ => nil).
-End cse.
-
-Definition example_expr_simplified := Eval vm_compute in Eta.ExprEta (InlineConst is_const (ANormal example_expr)).
-Compute CSE example_expr_simplified.
-
-Definition example_expr_compiled
- := Eval vm_compute in
- match Named.Compile.compile (example_expr_simplified _) (List.map Pos.of_nat (seq 1 20)) as v return match v with Some _ => _ | _ => _ end with
- | Some v => v
- | None => True
- end.
-
-Compute register_reassign (InContext:=PositiveContext_nd) (ReverseContext:=PositiveContext_nd) Pos.eqb empty empty example_expr_compiled (Some 1%positive :: Some 2%positive :: None :: List.map (@Some _) (List.map Pos.of_nat (seq 3 20))).
-
-Module bounds.
- Record bounded := { lower : nat ; value : nat ; upper : nat }.
- Definition map_bounded_f2 (f : nat -> nat -> nat) (swap_on_arg2 : bool) (x y : bounded)
- := {| lower := f (lower x) (if swap_on_arg2 then upper y else lower y);
- value := f (value x) (value y);
- upper := f (upper x) (if swap_on_arg2 then lower y else upper y) |}.
- Definition bounded_pf := { b : bounded | lower b <= value b <= upper b }.
- Definition add_bounded_pf (x y : bounded_pf) : bounded_pf.
- Proof.
- exists (map_bounded_f2 plus false (proj1_sig x) (proj1_sig y)).
- simpl; let x := x in let y := y in abstract (destruct x, y; simpl; omega).
- Defined.
- Definition mul_bounded_pf (x y : bounded_pf) : bounded_pf.
- Proof.
- exists (map_bounded_f2 mult false (proj1_sig x) (proj1_sig y)).
- simpl; let x := x in let y := y in abstract (destruct x, y; simpl; nia).
- Defined.
- Definition sub_bounded_pf (x y : bounded_pf) : bounded_pf.
- Proof.
- exists (map_bounded_f2 minus true (proj1_sig x) (proj1_sig y)).
- simpl; let x := x in let y := y in abstract (destruct x, y; simpl; omega).
- Defined.
- Definition interp_base_type_bounds (v : base_type) : Type :=
- match v with
- | Tnat => { b : bounded | lower b <= value b <= upper b }
- end.
- Definition constant_bounded t (x : interp_base_type t) : interp_base_type_bounds t.
- Proof.
- destruct t.
- exists {| lower := x ; value := x ; upper := x |}.
- simpl; split; reflexivity.
- Defined.
- Definition interp_op_bounds src dst (f : op src dst) : interp_flat_type interp_base_type_bounds src -> interp_flat_type interp_base_type_bounds dst
- := match f with
- | Const v => fun _ => constant_bounded Tnat v
- | Add => fun xy => add_bounded_pf (fst xy) (snd xy)
- | Mul => fun xy => mul_bounded_pf (fst xy) (snd xy)
- | Sub => fun xy => sub_bounded_pf (fst xy) (snd xy)
- end%nat.
- Fixpoint constant_bounds t
- : interp_flat_type interp_base_type t -> interp_flat_type interp_base_type_bounds t
- := match t with
- | Tbase t => constant_bounded t
- | Unit => fun _ => tt
- | Prod _ _ => fun x => (constant_bounds _ (fst x), constant_bounds _ (snd x))
- end.
-
- Compute example_expr_simplified.
- Local Open Scope zrange_scope.
- Eval compute in
- (fun x (xpf : 0 <= x <= 10) y (ypf : 100 <= y <= 1000)
- => let (l, _, u) :=
- proj1_sig
- (Syntax.Interp
- interp_op_bounds example_expr
- (exist _ {| lower := 0 ; value := x ; upper := 10 |} xpf,
- exist _ {| lower := 100 ; value := y ; upper := 1000 |} ypf))
- in {| ZRange.lower := Z.of_nat l ; ZRange.upper := Z.of_nat u |}).
-End bounds.
diff --git a/src/Compilers/Tuple.v b/src/Compilers/Tuple.v
deleted file mode 100644
index 8e47bf012..000000000
--- a/src/Compilers/Tuple.v
+++ /dev/null
@@ -1,88 +0,0 @@
-Require Import Crypto.Util.Tuple.
-Require Import Crypto.Compilers.Syntax.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
-
- Section interp.
- Section flat_type.
- Context {interp_base_type : base_type_code -> Type}.
- Local Notation interp_flat_type := (interp_flat_type interp_base_type).
-
- Fixpoint flat_interp_tuple' {T n} : interp_flat_type (tuple' T n) -> Tuple.tuple' (interp_flat_type T) n
- := match n return interp_flat_type (tuple' T n) -> Tuple.tuple' (interp_flat_type T) n with
- | O => fun x => x
- | S n' => fun '((x, y) : interp_flat_type (tuple' T n' * T))
- => (@flat_interp_tuple' _ n' x, y)
- end.
- Definition flat_interp_tuple {T n} : interp_flat_type (tuple T n) -> Tuple.tuple (interp_flat_type T) n
- := match n return interp_flat_type (tuple T n) -> Tuple.tuple (interp_flat_type T) n with
- | O => fun x => x
- | S n' => @flat_interp_tuple' T n'
- end.
- Fixpoint flat_interp_untuple' {T n} : Tuple.tuple' (interp_flat_type T) n -> interp_flat_type (tuple' T n)
- := match n return Tuple.tuple' (interp_flat_type T) n -> interp_flat_type (tuple' T n) with
- | O => fun x => x
- | S n' => fun '((x, y) : Tuple.tuple' _ n' * _)
- => (@flat_interp_untuple' _ n' x, y)
- end.
- Definition flat_interp_untuple {T n} : Tuple.tuple (interp_flat_type T) n -> interp_flat_type (tuple T n)
- := match n return Tuple.tuple (interp_flat_type T) n -> interp_flat_type (tuple T n) with
- | O => fun x => x
- | S n' => @flat_interp_untuple' T n'
- end.
- Lemma flat_interp_untuple'_tuple' {T n v}
- : @flat_interp_untuple' T n (flat_interp_tuple' v) = v.
- Proof using Type. induction n; [ reflexivity | simpl; destruct v; rewrite IHn; reflexivity ]. Qed.
- Lemma flat_interp_untuple_tuple {T n v}
- : flat_interp_untuple (@flat_interp_tuple T n v) = v.
- Proof using Type. destruct n; [ reflexivity | apply flat_interp_untuple'_tuple' ]. Qed.
- Lemma flat_interp_tuple'_untuple' {T n v}
- : @flat_interp_tuple' T n (flat_interp_untuple' v) = v.
- Proof using Type. induction n; [ reflexivity | simpl; destruct v; rewrite IHn; reflexivity ]. Qed.
- Lemma flat_interp_tuple_untuple {T n v}
- : @flat_interp_tuple T n (flat_interp_untuple v) = v.
- Proof using Type. destruct n; [ reflexivity | apply flat_interp_tuple'_untuple' ]. Qed.
- End flat_type.
- End interp.
-
- Section interp2.
- Section flat_type.
- Context {interp_base_type1 interp_base_type2 : base_type_code -> Type}.
- Local Notation interp_flat_type1 := (interp_flat_type interp_base_type1).
- Local Notation interp_flat_type2 := (interp_flat_type interp_base_type2).
-
- Definition tuple_map {A B n} (f : interp_flat_type1 A -> interp_flat_type2 B) (v : interp_flat_type1 (tuple A n))
- : interp_flat_type2 (tuple B n)
- := flat_interp_untuple (Tuple.map f (flat_interp_tuple v)).
- End flat_type.
- End interp2.
-End language.
-Global Arguments flat_interp_tuple' {_ _ _ _} _.
-Global Arguments flat_interp_tuple {_ _ _ _} _.
-Global Arguments flat_interp_untuple' {_ _ _ _} _.
-Global Arguments flat_interp_untuple {_ _ _ _} _.
-Global Arguments tuple_map {_ _ _ _ _ n} _ _.
-
-Ltac unfold_flat_interp_tuple _ :=
- let handle n :=
- ltac:(let n' := (eval cbv in n) in
- progress change n with n') in
- repeat match goal with
- | [ |- context[@flat_interp_tuple _ _ _ ?n] ]
- => handle n
- | [ |- context[@flat_interp_tuple' _ _ _ ?n] ]
- => handle n
- | [ |- context[@flat_interp_untuple _ _ _ ?n] ]
- => handle n
- | [ |- context[@flat_interp_untuple' _ _ _ ?n] ]
- => handle n
- | [ |- context[@tuple _ _ ?n] ]
- => handle n
- | [ |- context[@tuple' _ _ ?n] ]
- => handle n
- end;
- cbv [flat_interp_tuple flat_interp_tuple' flat_interp_untuple flat_interp_untuple' tuple tuple'].
diff --git a/src/Compilers/TypeInversion.v b/src/Compilers/TypeInversion.v
deleted file mode 100644
index ab00a1dc5..000000000
--- a/src/Compilers/TypeInversion.v
+++ /dev/null
@@ -1,193 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Util.FixCoqMistakes.
-
-Section language.
- Context {base_type_code : Type}.
-
- Section flat.
- Context (P : flat_type base_type_code -> Type).
-
- Local Ltac t :=
- let H := fresh in
- intro H; intros;
- match goal with
- | [ p : _ |- _ ] => specialize (H _ p)
- end;
- cbv beta iota in *;
- try specialize (H eq_refl); simpl in *;
- try assumption.
-
- Definition preinvert_Tbase (Q : forall t, P (Tbase t) -> Type)
- : (forall t (p : P t), match t return P t -> Type with Tbase _ => Q _ | _ => fun _ => True end p)
- -> forall t p, Q t p.
- Proof. t. Defined.
-
- Definition preinvert_Unit (Q : P Unit -> Type)
- : (forall t (p : P t), match t return P t -> Type with Unit => Q | _ => fun _ => True end p)
- -> forall p, Q p.
- Proof. t. Defined.
-
- Definition preinvert_Prod (Q : forall A B, P (Prod A B) -> Type)
- : (forall t (p : P t), match t return P t -> Type with Prod _ _ => Q _ _ | _ => fun _ => True end p)
- -> forall A B p, Q A B p.
- Proof. t. Defined.
-
- Definition preinvert_Prod2 (Q : forall A B, P (Prod (Tbase A) (Tbase B)) -> Type)
- : (forall t (p : P t), match t return P t -> Type with Prod (Tbase _) (Tbase _) => Q _ _ | _ => fun _ => True end p)
- -> forall A B p, Q A B p.
- Proof. t. Defined.
-
- Definition preinvert_Prod2_same (Q : forall A, P (Prod (Tbase A) (Tbase A)) -> Type)
- : (forall t (p : P t), match t return P t -> Type with
- | Prod (Tbase A) (Tbase B)
- => fun p => forall pf : A = B, Q B (eq_rect _ (fun a => P (Prod (Tbase a) (Tbase B))) p _ pf)
- | _ => fun _ => True
- end p)
- -> forall A p, Q A p.
- Proof. t. Defined.
-
- Definition preinvert_Prod3 (Q : forall A B C, P (Tbase A * Tbase B * Tbase C)%ctype -> Type)
- : (forall t (p : P t), match t return P t -> Type with Prod (Prod (Tbase _) (Tbase _)) (Tbase _) => Q _ _ _ | _ => fun _ => True end p)
- -> forall A B C p, Q A B C p.
- Proof. t. Defined.
-
- Definition preinvert_Prod4 (Q : forall A B C D, P (Tbase A * Tbase B * Tbase C * Tbase D)%ctype -> Type)
- : (forall t (p : P t), match t return P t -> Type with Prod (Prod (Prod (Tbase _) (Tbase _)) (Tbase _)) (Tbase _) => Q _ _ _ _ | _ => fun _ => True end p)
- -> forall A B C D p, Q A B C D p.
- Proof. t. Defined.
- End flat.
-
- Definition preinvert_Arrow (P : type base_type_code -> Type) (Q : forall A B, P (Arrow A B) -> Type)
- : (forall t (p : P t), match t return P t -> Type with
- | Arrow A B => Q A B
- end p)
- -> forall A B p, Q A B p.
- Proof.
- intros H A B p; specialize (H _ p); assumption.
- Defined.
-
- Section encode_decode.
- Definition flat_type_code (t1 t2 : flat_type base_type_code) : Prop
- := match t1, t2 with
- | Unit, Unit => True
- | Tbase t1, Tbase t2 => t1 = t2
- | Prod A B, Prod A' B' => A = A' /\ B = B'
- | Unit, _
- | Tbase _, _
- | Prod _ _, _
- => False
- end.
-
- Definition type_code (t1 t2 : type base_type_code) : Prop
- := domain t1 = domain t2 /\ codomain t1 = codomain t2.
-
- Definition flat_type_encode (x y : flat_type base_type_code) : x = y -> flat_type_code x y.
- Proof. intro p; destruct p, x; repeat constructor. Defined.
- Definition type_encode (x y : type base_type_code) : x = y -> type_code x y.
- Proof. intro p; destruct p, x; repeat constructor. Defined.
-
- Definition flat_type_decode (x y : flat_type base_type_code) : flat_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 type_decode (x y : type base_type_code) : type_code x y -> x = y.
- Proof.
- destruct x, y; simpl; intro H;
- try first [ exfalso; assumption
- | apply f_equal; assumption
- | apply f_equal2; destruct H; assumption ].
- Defined.
- Definition path_flat_type_rect {x y : flat_type base_type_code} (Q : x = y -> Type)
- (f : forall p, Q (flat_type_decode x y p))
- : forall p, Q p.
- Proof. intro p; specialize (f (flat_type_encode x y p)); destruct x, p; exact f. Defined.
- Definition path_type_rect {x y : type base_type_code} (Q : x = y -> Type)
- (f : forall p, Q (type_decode x y p))
- : forall p, Q p.
- Proof. intro p; specialize (f (type_encode x y p)); destruct x, p; exact f. Defined.
- End encode_decode.
-End language.
-
-Ltac preinvert_one_type e :=
- lazymatch type of e with
- | ?P (Tbase ?T)
- => is_var T;
- move e at top;
- revert dependent T;
- refine (preinvert_Tbase P _ _)
- | ?P (Prod (Tbase ?A) (Tbase ?A))
- => is_var A;
- move e at top; revert dependent A;
- refine (preinvert_Prod2_same P _ _)
- | ?P (Prod (Tbase ?A) (Tbase ?B))
- => is_var A; is_var B;
- move e at top; revert dependent A; intros A e;
- move e at top; revert dependent B; revert A;
- refine (preinvert_Prod2 P _ _)
- | ?P (Prod (Prod (Tbase ?A) (Tbase ?B)) (Tbase ?C))
- => is_var A; is_var B; is_var C;
- move e at top; revert dependent A; intros A e;
- move e at top; revert dependent B; intros B e;
- move e at top; revert dependent C; revert A B;
- refine (preinvert_Prod3 P _ _)
- | ?P (Prod (Prod (Prod (Tbase ?A) (Tbase ?B)) (Tbase ?C)) (Tbase ?D))
- => is_var A; is_var B; is_var C; is_var D;
- move e at top; revert dependent A; intros A e;
- move e at top; revert dependent B; intros B e;
- move e at top; revert dependent C; intros C e;
- move e at top; revert dependent D; revert A B C;
- refine (preinvert_Prod4 P _ _)
- | ?P (Prod ?A ?B)
- => is_var A; is_var B;
- move e at top; revert dependent A; intros A e;
- move e at top; revert dependent B; revert A;
- refine (preinvert_Prod P _ _)
- | ?P Unit
- => revert dependent e;
- refine (preinvert_Unit P _ _)
- | ?P (Arrow ?A ?B)
- => is_var A; is_var B;
- move e at top; revert dependent A; intros A e;
- move e at top; revert dependent B; revert A;
- refine (preinvert_Arrow P _ _)
- end.
-
-Ltac induction_type_in_using H rect :=
- induction H as [H] using (rect _ _ _);
- cbv [flat_type_code 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
- | _ /\ _ => destruct H as [H1 H2]
- end.
-Ltac inversion_flat_type_step :=
- lazymatch goal with
- | [ H : _ = Tbase _ |- _ ]
- => induction_type_in_using H @path_flat_type_rect
- | [ H : Tbase _ = _ |- _ ]
- => induction_type_in_using H @path_flat_type_rect
- | [ H : _ = Prod _ _ |- _ ]
- => induction_type_in_using H @path_flat_type_rect
- | [ H : Prod _ _ = _ |- _ ]
- => induction_type_in_using H @path_flat_type_rect
- | [ H : _ = Unit |- _ ]
- => induction_type_in_using H @path_flat_type_rect
- | [ H : Unit = _ |- _ ]
- => induction_type_in_using H @path_flat_type_rect
- end.
-Ltac inversion_flat_type := repeat inversion_flat_type_step.
-
-Ltac inversion_type_step :=
- lazymatch goal with
- | [ H : _ = Arrow _ _ |- _ ]
- => induction_type_in_using H @path_type_rect
- | [ H : Arrow _ _ = _ |- _ ]
- => induction_type_in_using H @path_type_rect
- end.
-Ltac inversion_type := repeat inversion_type_step.
diff --git a/src/Compilers/TypeUtil.v b/src/Compilers/TypeUtil.v
deleted file mode 100644
index 050374562..000000000
--- a/src/Compilers/TypeUtil.v
+++ /dev/null
@@ -1,35 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Util.Notations.
-
-Local Open Scope expr_scope.
-
-Section language.
- Context {base_type_code : Type}
- (base_type_beq : base_type_code -> base_type_code -> bool)
- (base_type_leb : base_type_code -> base_type_code -> bool).
- Local Infix "<=?" := base_type_leb : expr_scope.
- Local Infix "=?" := base_type_beq : expr_scope.
-
- Definition base_type_min (a b : base_type_code) : base_type_code
- := if a <=? b then a else b.
- Definition base_type_max (a b : base_type_code) : base_type_code
- := if a <=? b then b else a.
- Section gen.
- Context (join : base_type_code -> base_type_code -> base_type_code).
- Fixpoint flat_type_join {t : flat_type base_type_code}
- : interp_flat_type (fun _ => base_type_code) t -> option base_type_code
- := match t with
- | Tbase _ => fun v => Some v
- | Unit => fun _ => None
- | Prod A B
- => fun v => match @flat_type_join A (fst v), @flat_type_join B (snd v) with
- | Some a, Some b => Some (join a b)
- | Some a, None => Some a
- | None, Some b => Some b
- | None, None => None
- end
- end.
- End gen.
- Definition flat_type_min {t} := @flat_type_join base_type_min t.
- Definition flat_type_max {t} := @flat_type_join base_type_max t.
-End language.
diff --git a/src/Compilers/Wf.v b/src/Compilers/Wf.v
deleted file mode 100644
index 85c24886b..000000000
--- a/src/Compilers/Wf.v
+++ /dev/null
@@ -1,68 +0,0 @@
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Util.Notations.
-
-Create HintDb wf discriminated.
-
-Ltac solve_wf_side_condition := solve [ eassumption | eauto 250 with wf ].
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}.
-
- Local Notation eP2 := (fun t1t2 => var1 (fst t1t2) * var2 (snd t1t2))%type (only parsing).
- Local Notation eP := (fun t => var1 t * var2 t)%type (only parsing).
- Local Notation "x == y" := (existT eP _ (x, y)%core).
- Fixpoint flatten_binding_list2 {t1 t2} (x : interp_flat_type var1 t1) (y : interp_flat_type var2 t2) : list (sigT eP2)
- := (match t1, t2 return interp_flat_type var1 t1 -> interp_flat_type var2 t2 -> list _ with
- | Tbase t1, Tbase t2 => fun x y => existT eP2 (t1, t2)%core (x, y)%core :: nil
- | Unit, Unit => fun x y => nil
- | Prod t0 t1, Prod t0' t1'
- => fun x y => @flatten_binding_list2 _ _ (snd x) (snd y) ++ @flatten_binding_list2 _ _ (fst x) (fst y)
- | Tbase _, _
- | Unit, _
- | Prod _ _, _
- => fun _ _ => nil
- end x y)%list.
- Fixpoint flatten_binding_list {t} (x : interp_flat_type var1 t) (y : interp_flat_type var2 t) : list (sigT eP)
- := (match t return interp_flat_type var1 t -> interp_flat_type var2 t -> list _ with
- | Tbase _ => fun x y => (x == y) :: nil
- | Unit => fun x y => nil
- | Prod t0 t1 => fun x y => @flatten_binding_list _ (snd x) (snd y) ++ @flatten_binding_list _ (fst x) (fst y)
- end x y)%list.
-
- Inductive wff : list (sigT eP) -> forall {t}, @exprf var1 t -> @exprf var2 t -> Prop :=
- | WfTT : forall G, @wff G _ TT TT
- | WfVar : forall G (t : base_type_code) x x', List.In (x == x') G -> @wff G (Tbase t) (Var x) (Var x')
- | WfOp : forall G {t} {tR} (e : @exprf var1 t) (e' : @exprf var2 t) op,
- wff G e e'
- -> wff G (Op (tR := tR) op e) (Op (tR := tR) op e')
- | WfLetIn : forall G t1 t2 e1 e1' (e2 : interp_flat_type var1 t1 -> @exprf var1 t2) e2',
- wff G e1 e1'
- -> (forall x1 x2, wff (flatten_binding_list x1 x2 ++ G) (e2 x1) (e2' x2))
- -> wff G (LetIn e1 e2) (LetIn e1' e2')
- | WfPair : forall G {t1} {t2} (e1: @exprf var1 t1) (e2: @exprf var1 t2)
- (e1': @exprf var2 t1) (e2': @exprf var2 t2),
- wff G e1 e1'
- -> wff G e2 e2'
- -> wff G (Pair e1 e2) (Pair e1' e2').
- Inductive wf : forall {t}, @expr var1 t -> @expr var2 t -> Prop :=
- | WfAbs : forall A B e e',
- (forall x x', @wff (flatten_binding_list x x') B (e x) (e' x'))
- -> @wf (Arrow A B) (Abs e) (Abs e').
- End with_var.
-
- Definition Wf {t} (E : @Expr t) := forall var1 var2, wf (E var1) (E var2).
-End language.
-
-Global Arguments wff {_ _ _ _} G {t} _ _.
-Global Arguments wf {_ _ _ _ t} _ _.
-Global Arguments Wf {_ _ t} _.
-
-Hint Constructors wf wff : wf.
diff --git a/src/Compilers/WfInversion.v b/src/Compilers/WfInversion.v
deleted file mode 100644
index bf8c93ade..000000000
--- a/src/Compilers/WfInversion.v
+++ /dev/null
@@ -1,208 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Util.Sigma.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Equality.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Notations.
-
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
- Local Notation wf := (@wf base_type_code op).
- Local Notation Wf := (@Wf base_type_code op).
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}.
-
- Local Notation eP := (fun t => var1 t * var2 t)%type (only parsing).
- Local Notation "x == y" := (existT eP _ (x, y)).
-
- Definition wff_code (G : list (sigT eP)) {t} (e1 : @exprf var1 t) : forall (e2 : @exprf var2 t), Prop
- := match e1 in Syntax.exprf _ _ t return exprf t -> Prop with
- | TT
- => fun e2
- => TT = e2
- | Var t v1
- => fun e2
- => match invert_Var e2 with
- | Some v2 => List.In (v1 == v2) G
- | None => False
- end
- | Op t1 tR opc1 args1
- => fun e2
- => match invert_Op e2 with
- | Some (existT t2 (opc2, args2))
- => { pf : t1 = t2
- | eq_rect _ (fun t => op t tR) opc1 _ pf = opc2
- /\ wff G (eq_rect _ exprf args1 _ pf) args2 }
- | None => False
- end
- | LetIn tx1 ex1 tC1 eC1
- => fun e2
- => match invert_LetIn e2 with
- | Some (existT tx2 (ex2, eC2))
- => { pf : tx1 = tx2
- | wff G (eq_rect _ exprf ex1 _ pf) ex2
- /\ (forall x1 x2,
- wff (flatten_binding_list x1 x2 ++ G)%list
- (eC1 x1) (eC2 (eq_rect _ _ x2 _ pf))) }
- | None => False
- end
- | Pair tx1 ex1 ty1 ey1
- => fun e2
- => match invert_Pair e2 with
- | Some (ex2, ey2) => wff G ex1 ex2 /\ wff G ey1 ey2
- | None => False
- end
- end.
-
- Local Ltac t :=
- repeat match goal with
- | _ => progress simpl in *
- | _ => progress subst
- | _ => progress inversion_option
- | _ => progress invert_expr_subst
- | [ H : Some _ = _ |- _ ] => symmetry in H
- | _ => assumption
- | _ => reflexivity
- | _ => constructor
- | _ => progress destruct_head False
- | _ => progress destruct_head and
- | _ => progress destruct_head sig
- | _ => progress break_match_hyps
- | _ => progress break_match
- | [ |- and _ _ ] => split
- | _ => exists eq_refl
- | _ => intro
- | [ e : expr (Arrow _ _) |- _ ]
- => let H := fresh in
- let f := fresh in
- remember (invert_Abs e) as f eqn:H;
- symmetry in H;
- apply invert_Abs_Some in H
- end.
-
- Definition wff_encode {G t e1 e2} (v : @wff var1 var2 G t e1 e2) : @wff_code G t e1 e2.
- Proof.
- destruct v; t.
- Defined.
-
- Definition wff_decode {G t e1 e2} (v : @wff_code G t e1 e2) : @wff var1 var2 G t e1 e2.
- Proof.
- destruct e1; t.
- Defined.
-
- Definition wff_endecode {G t e1 e2} v : @wff_decode G t e1 e2 (@wff_encode G t e1 e2 v) = v.
- Proof using Type.
- destruct v; reflexivity.
- Qed.
-
- Definition wff_deencode {G t e1 e2} v : @wff_encode G t e1 e2 (@wff_decode G t e1 e2 v) = v.
- Proof using Type.
- destruct e1; simpl in *;
- move e2 at top;
- lazymatch type of e2 with
- | exprf Unit
- => subst; reflexivity
- | exprf (Tbase ?t)
- => revert dependent t;
- intros ? e2
- | exprf (Prod ?A ?B)
- => revert dependent A;
- intros ? e2;
- move e2 at top;
- revert dependent B;
- intros ? e2
- | exprf ?t
- => revert dependent t;
- intros ? e2
- end;
- refine match e2 with
- | TT => _
- | _ => _
- end;
- t.
- Qed.
-
- Definition wf_code {t} (e1 : @expr var1 t) : forall (e2 : @expr var2 t), Prop
- := match e1 in Syntax.expr _ _ t return expr t -> Prop with
- | Abs src dst f1
- => fun e2
- => let f2 := invert_Abs e2 in
- forall (x : interp_flat_type var1 src) (x' : interp_flat_type var2 src),
- wff (flatten_binding_list x x') (f1 x) (f2 x')
- end.
-
- Definition wf_encode {t e1 e2} (v : @wf var1 var2 t e1 e2) : @wf_code t e1 e2.
- Proof.
- destruct v; t.
- Defined.
-
- Definition wf_decode {t e1 e2} (v : @wf_code t e1 e2) : @wf var1 var2 t e1 e2.
- Proof.
- destruct e1; t.
- Defined.
-
- Definition wf_endecode {t e1 e2} v : @wf_decode t e1 e2 (@wf_encode t e1 e2 v) = v.
- Proof using Type.
- destruct v; reflexivity.
- Qed.
-
- Definition wf_deencode {t e1 e2} v : @wf_encode t e1 e2 (@wf_decode t e1 e2 v) = v.
- Proof using Type.
- destruct e1 as [src dst f1].
- revert dependent f1.
- refine match e2 with
- | Abs _ _ f2 => _
- end.
- reflexivity.
- Qed.
- End with_var.
-End language.
-
-Ltac is_expr_constructor arg :=
- lazymatch arg with
- | Op _ _ => idtac
- | TT => idtac
- | Var _ => idtac
- | LetIn _ _ => idtac
- | Pair _ _ => idtac
- | Abs _ => idtac
- end.
-
-Ltac inversion_wf_step_gen guard_tac :=
- let postprocess H :=
- (cbv [wff_code wf_code] in H;
- simpl in H;
- try match type of H with
- | True => clear H
- | False => exfalso; exact H
- end) in
- match goal with
- | [ H : wff _ ?x ?y |- _ ]
- => guard_tac x y;
- apply wff_encode in H; postprocess H
- | [ H : wf ?x ?y |- _ ]
- => guard_tac x y;
- apply wf_encode in H; postprocess H
- end.
-Ltac inversion_wf_step_constr :=
- inversion_wf_step_gen ltac:(fun x y => is_expr_constructor x; is_expr_constructor y).
-Ltac inversion_wf_step_one_constr :=
- inversion_wf_step_gen ltac:(fun x y => first [ is_expr_constructor x | is_expr_constructor y]).
-Ltac inversion_wf_step_var :=
- inversion_wf_step_gen ltac:(fun x y => first [ is_var x; is_var y; fail 1 | idtac ]).
-Ltac inversion_wf_step := first [ inversion_wf_step_constr | inversion_wf_step_var ].
-Ltac inversion_wf_constr := repeat inversion_wf_step_constr.
-Ltac inversion_wf_one_constr := repeat inversion_wf_step_one_constr.
-Ltac inversion_wf := repeat inversion_wf_step.
diff --git a/src/Compilers/WfProofs.v b/src/Compilers/WfProofs.v
deleted file mode 100644
index 4a94f2591..000000000
--- a/src/Compilers/WfProofs.v
+++ /dev/null
@@ -1,473 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.WfInversion.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Util.Sigma Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.Tactics.SplitInContext.
-Require Import Crypto.Util.Tactics.BreakMatch.
-
-Local Open Scope ctype_scope.
-Section language.
- Context {base_type_code : Type}
- {op : flat_type base_type_code -> flat_type base_type_code -> Type}.
-
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation Expr := (@Expr base_type_code op).
- Local Notation wff := (@wff base_type_code op).
-
- Section with_var.
- Context {var1 var2 : base_type_code -> Type}.
- Local Hint Constructors Wf.wff.
-
- Lemma wff_app' {g G0 G1 t e1 e2}
- (wf : @wff var1 var2 (G0 ++ G1) t e1 e2)
- : wff (G0 ++ g ++ G1) e1 e2.
- Proof using Type.
- rewrite !List.app_assoc.
- revert wf; remember (G0 ++ G1)%list as G eqn:?; intro wf.
- revert dependent G0. revert dependent G1.
- induction wf; simpl in *; constructor; simpl; eauto.
- { subst; rewrite !List.in_app_iff in *; intuition. }
- { intros; subst.
- rewrite !List.app_assoc; eauto using List.app_assoc. }
- Qed.
-
- Lemma wff_app_pre {g G t e1 e2}
- (wf : @wff var1 var2 G t e1 e2)
- : wff (g ++ G) e1 e2.
- Proof using Type.
- apply (@wff_app' _ nil); assumption.
- Qed.
-
- Lemma wff_app_post {g G t e1 e2}
- (wf : @wff var1 var2 G t e1 e2)
- : wff (G ++ g) e1 e2.
- Proof using Type.
- pose proof (@wff_app' g G nil t e1 e2) as H.
- rewrite !List.app_nil_r in *; auto.
- Qed.
-
- Lemma wff_in_impl_Proper G0 G1 {t} e1 e2
- : @wff var1 var2 G0 t e1 e2
- -> (forall x, List.In x G0 -> List.In x G1)
- -> @wff var1 var2 G1 t e1 e2.
- Proof using Type.
- intro wf; revert G1; induction wf;
- repeat match goal with
- | _ => setoid_rewrite List.in_app_iff
- | _ => progress intros
- | _ => progress simpl in *
- | [ |- wff _ _ _ ] => constructor
- | [ H : _ |- _ ] => apply H
- | _ => solve [ intuition eauto ]
- end.
- Qed.
-
- Local Hint Resolve List.in_app_or List.in_or_app.
- Local Hint Extern 1 => progress unfold List.In in *.
- Local Hint Resolve wff_in_impl_Proper.
-
- Lemma wff_SmartVarf {t} x1 x2
- : @wff var1 var2 (flatten_binding_list x1 x2) t (SmartVarf x1) (SmartVarf x2).
- Proof using Type.
- unfold SmartVarf.
- induction t; simpl; constructor; eauto.
- Qed.
-
- Local Hint Resolve wff_SmartVarf.
-
- Lemma wff_SmartVarVarf G {t t'} v1 v2 x1 x2
- (Hin : List.In (existT (fun t : base_type_code => (exprf (Tbase t) * exprf (Tbase t))%type) t (x1, x2))
- (flatten_binding_list (SmartVarVarf v1) (SmartVarVarf v2)))
- : @wff var1 var2 (flatten_binding_list (t:=t') v1 v2 ++ G) (Tbase t) x1 x2.
- Proof using Type.
- revert dependent G; induction t'; intros; simpl in *; try tauto.
- { intuition (inversion_sigma; inversion_prod; subst; simpl; eauto).
- constructor; eauto. }
- { unfold SmartVarVarf in *; simpl in *.
- apply List.in_app_iff in Hin.
- intuition (inversion_sigma; inversion_prod; subst; eauto).
- { rewrite <- !List.app_assoc; eauto. } }
- Qed.
-
- Lemma wff_SmartVarVarf_nil {t t'} v1 v2 x1 x2
- (Hin : List.In (existT (fun t : base_type_code => (exprf (Tbase t) * exprf (Tbase t))%type) t (x1, x2))
- (flatten_binding_list (SmartVarVarf v1) (SmartVarVarf v2)))
- : @wff var1 var2 (flatten_binding_list (t:=t') v1 v2) (Tbase t) x1 x2.
- Proof using Type.
- apply wff_SmartVarVarf with (G:=nil) in Hin.
- rewrite List.app_nil_r in Hin; assumption.
- Qed.
-
- Lemma In_G_wff_SmartVarf G t v1 v2 e
- (Hwf : @wff var1 var2 G t (SmartVarf v1) (SmartVarf v2))
- (Hin : List.In e (flatten_binding_list v1 v2))
- : List.In e G.
- Proof using Type.
- induction t;
- repeat match goal with
- | _ => assumption
- | [ H : False |- _ ] => exfalso; assumption
- | _ => progress subst
- | _ => progress destruct_head' and
- | [ H : context[List.In _ (_ ++ _)] |- _ ] => rewrite List.in_app_iff in H
- | [ H : context[SmartVarf _] |- _ ] => rewrite SmartVarf_Pair in H
- | _ => progress simpl in *
- | _ => progress destruct_head' or
- | _ => solve [ eauto with nocore ]
- | _ => progress inversion_wf
- end.
- Qed.
-
- Lemma wff_SmartPairf G {t t'} v1 v2 x1 x2
- (Hin : List.In (existT (fun t : base_type_code => (exprf (Tbase t) * exprf (Tbase t))%type) t (x1, x2))
- (flatten_binding_list (t:=t') v1 v2))
- (Hwf : wff G (SmartPairf v1) (SmartPairf v2))
- : @wff var1 var2 G (Tbase t) x1 x2.
- Proof using Type.
- revert dependent G; induction t'; intros; simpl in *; try tauto.
- { intuition (inversion_sigma; inversion_prod; subst; simpl; eauto). }
- { unfold SmartPairf in *; simpl in *.
- inversion_wf; destruct_head'_and.
- apply List.in_app_iff in Hin.
- intuition (inversion_sigma; inversion_prod; subst; eauto). }
- Qed.
-
- Section with_interp.
- Context {interp_base_type : base_type_code -> Type}.
-
- Lemma wff_invert_PairsConst G {t} e1 e2
- invert_Const1 invert_Const2
- (Hinvert_Const : forall s d (opv : op s d) G args1 args2,
- invert_PairsConst invert_Const1 args1
- = invert_PairsConst invert_Const2 args2
- -> wff G args1 args2
- -> invert_Const1 s d opv args1
- = invert_Const2 s d opv args2)
- (Hwf : wff G (t:=t) (var1:=var1) (var2:=var2) e1 e2)
- : invert_PairsConst (interp_base_type:=interp_base_type) invert_Const1 e1
- = invert_PairsConst invert_Const2 e2.
- Proof using Type.
- induction Hwf; simpl in *; break_innermost_match; try congruence; eauto.
- Qed.
- End with_interp.
-
- Lemma wff_SmartPairf_SmartVarfMap_same {var} G {t} v f g
- (Hfg : forall t v, wff G (f t v) (g t v))
- : wff G (t:=t) (var1:=var1) (var2:=var2)
- (SmartPairf (SmartVarfMap f v))
- (SmartPairf (SmartVarfMap (var:=var) g v)).
- Proof.
- induction t; try solve [ cbv [SmartPairf]; simpl; auto ].
- rewrite !SmartVarfMap_Pair, !SmartPairf_Pair; auto.
- Qed.
-
- Lemma wff_SmartPairf_SmartValf G {t} f g
- (Hfg : forall t, wff G (f t) (g t))
- : wff G (t:=t) (var1:=var1) (var2:=var2)
- (SmartPairf (SmartValf (fun t => exprf (Tbase t)) f _))
- (SmartPairf (SmartValf (fun t => exprf (Tbase t)) g _)).
- Proof.
- induction t; try solve [ cbv [SmartPairf]; simpl; auto ].
- Qed.
- End with_var.
-
- Section with_var2.
- Context {base_type_code2 : Type}
- {var1 : base_type_code -> Type}
- {var2 : base_type_code2 -> Type}.
- Local Hint Constructors Wf.wff.
-
- Lemma In_flatten_binding_list_untransfer_interp_flat_type
- var1' var2' f_base
- (f_var12 : forall t, var1 t -> var2 (f_base t))
- (f_var21 : forall t, var2 (f_base t) -> var1 t)
- (f_var'12 : forall t, var1' t -> var2' (f_base t))
- (f_var'21 : forall t, var2' (f_base t) -> var1' t)
- (Hvar12 : forall t v, f_var12 t (f_var21 t v) = v)
- (Hvar'12 : forall t v, f_var'12 t (f_var'21 t v) = v)
- : forall T t x x' x1 x2,
- List.In
- (existT _ t (x, x'))
- (flatten_binding_list
- (t:=T)
- (untransfer_interp_flat_type f_base f_var21 x1)
- (untransfer_interp_flat_type f_base f_var'21 x2))
- -> List.In
- (existT _ (f_base t) (f_var12 t x, f_var'12 t x'))
- (flatten_binding_list x1 x2).
- Proof.
- induction T;
- repeat first [ progress simpl in *
- | progress intros
- | progress subst
- | exfalso; assumption
- | progress inversion_sigma
- | progress inversion_prod
- | progress destruct_head'_or
- | rewrite List.in_app_iff
- | solve [ eauto ]
- | rewrite Hvar12, Hvar'12
- | match goal with
- | [ H : _ |- _ ] => rewrite List.in_app_iff in H
- end ].
- Qed.
- End with_var2.
-
- Definition duplicate_type {var1 var2}
- : { t : base_type_code & (var1 t * var2 t)%type }
- -> { t1t2 : _ & (var1 (fst t1t2) * var2 (snd t1t2))%type }
- := fun txy => existT _ (projT1 txy, projT1 txy) (projT2 txy).
- Definition duplicate_types {var1 var2}
- := List.map (@duplicate_type var1 var2).
-
- Lemma flatten_binding_list_flatten_binding_list2
- {var1 var2 t1} x1 x2
- : duplicate_types (@flatten_binding_list base_type_code var1 var2 t1 x1 x2)
- = @flatten_binding_list2 base_type_code var1 var2 t1 t1 x1 x2.
- Proof using Type.
- induction t1; simpl; try reflexivity.
- rewrite_hyp <- !*.
- unfold duplicate_types; rewrite List.map_app; reflexivity.
- Qed.
-
- Local Ltac flatten_t :=
- repeat first [ reflexivity
- | intro
- | progress simpl @flatten_binding_list
- | progress simpl @flatten_binding_list2
- | rewrite !List.map_app
- | progress simpl in *
- | rewrite_hyp <- !*; reflexivity
- | rewrite_hyp !*; reflexivity ].
-
- Lemma flatten_binding_list2_SmartVarfMap
- {var1 var1' var2 var2' t1 t2} f g (x1 : interp_flat_type var1 t1) (x2 : interp_flat_type var2 t2)
- : flatten_binding_list2 (var1:=var1') (var2:=var2') (base_type_code:=base_type_code) (SmartVarfMap f x1) (SmartVarfMap g x2)
- = List.map (fun txy => existT _ (projT1 txy) (f _ (fst (projT2 txy)), g _ (snd (projT2 txy)))%core)
- (flatten_binding_list2 x1 x2).
- Proof using Type.
- revert dependent t2; induction t1, t2; flatten_t.
- Qed.
-
- Lemma flatten_binding_list2_SmartVarfMap1
- {var1 var1' var2' t1 t2} f (x1 : interp_flat_type var1 t1) (x2 : interp_flat_type var2' t2)
- : flatten_binding_list2 (var1:=var1') (var2:=var2') (base_type_code:=base_type_code) (SmartVarfMap f x1) x2
- = List.map (fun txy => existT _ (projT1 txy) (f _ (fst (projT2 txy)), snd (projT2 txy))%core)
- (flatten_binding_list2 x1 x2).
- Proof using Type.
- revert dependent t2; induction t1, t2; flatten_t.
- Qed.
-
- Lemma flatten_binding_list2_SmartVarfMap2
- {var1' var2 var2' t1 t2} g (x1 : interp_flat_type var1' t1) (x2 : interp_flat_type var2 t2)
- : flatten_binding_list2 (var1:=var1') (var2:=var2') (base_type_code:=base_type_code) x1 (SmartVarfMap g x2)
- = List.map (fun txy => existT _ (projT1 txy) (fst (projT2 txy), g _ (snd (projT2 txy)))%core)
- (flatten_binding_list2 x1 x2).
- Proof using Type.
- revert dependent t2; induction t1, t2; flatten_t.
- Qed.
-
- Lemma flatten_binding_list_SmartVarfMap
- {var1 var1' var2 var2' t} f g (x1 : interp_flat_type var1 t) (x2 : interp_flat_type var2 t)
- : flatten_binding_list (var1:=var1') (var2:=var2') (base_type_code:=base_type_code) (SmartVarfMap f x1) (SmartVarfMap g x2)
- = List.map (fun txy => existT _ (projT1 txy) (f _ (fst (projT2 txy)), g _ (snd (projT2 txy)))%core)
- (flatten_binding_list x1 x2).
- Proof using Type. induction t; flatten_t. Qed.
-
- Lemma flatten_binding_list2_SmartValf
- {T1 T2} f g t1 t2
- : flatten_binding_list2 (base_type_code:=base_type_code) (SmartValf T1 f t1) (SmartValf T2 g t2)
- = List.map (fun txy => existT _ (projT1 txy) (f _, g _)%core)
- (flatten_binding_list2 (SmartFlatTypeUnMap t1) (SmartFlatTypeUnMap t2)).
- Proof using Type.
- revert dependent t2; induction t1, t2; flatten_t.
- Qed.
-
- Lemma flatten_binding_list_SmartValf
- {T1 T2} f g t
- : flatten_binding_list (base_type_code:=base_type_code) (SmartValf T1 f t) (SmartValf T2 g t)
- = List.map (fun txy => existT _ (projT1 txy) (f _, g _)%core)
- (flatten_binding_list (SmartFlatTypeUnMap t) (SmartFlatTypeUnMap t)).
- Proof using Type. induction t; flatten_t. Qed.
-
- Lemma flatten_binding_list_In_eq_iff
- {var} T x y
- : (forall t a b, List.In (existT _ t (a, b)) (@flatten_binding_list base_type_code var var T x y) -> a = b)
- <-> x = y.
- Proof using Type.
- induction T;
- repeat first [ exfalso; assumption
- | progress subst
- | progress inversion_sigma
- | progress inversion_prod
- | progress destruct_head' unit
- | progress destruct_head' prod
- | split
- | progress simpl in *
- | intro
- | progress destruct_head or
- | apply (f_equal2 (@pair _ _))
- | progress split_iff
- | solve [ auto using List.in_or_app ]
- | match goal with
- | [ H : List.In _ (_ ++ _) |- _ ] => rewrite List.in_app_iff in H
- | [ H : forall x y, x = y -> forall t a b, List.In _ _ -> _, H' : List.In _ _ |- _ ]
- => specialize (H _ _ eq_refl _ _ _ H')
- end ].
- Qed.
-
- Lemma flatten_binding_list_same_in_eq
- {var} {T x t a b}
- : List.In (existT _ t (a, b)) (@flatten_binding_list base_type_code var var T x x) -> a = b.
- Proof using Type. intro; eapply flatten_binding_list_In_eq_iff; eauto. Qed.
-
- Lemma flatten_binding_list_interpf_SmartPairf_same
- interp_base_type
- (interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d)
- T T' x y e
- (Hin : List.In
- (existT
- (fun t => (exprf (Tbase t) * interp_base_type t)%type)
- T (x, y))
- (flatten_binding_list (t:=T') e (interpf interp_op (SmartMap.SmartPairf e))))
- : interpf interp_op x = y.
- Proof using Type.
- induction T';
- repeat first [ progress simpl in *
- | reflexivity
- | exfalso; assumption
- | progress subst
- | progress inversion_sigma
- | progress inversion_prod
- | progress destruct_head'_or
- | match goal with
- | [ H : List.In _ (_ ++ _) |- _ ] => rewrite List.in_app_iff in H
- | [ H : forall x y, x = y -> forall t a b, List.In _ _ -> _, H' : List.In _ _ |- _ ]
- => specialize (H _ _ eq_refl _ _ _ H')
- end
- | solve [ eauto ] ].
- Qed.
-
- Lemma flatten_binding_list_SmartVarfMap2_pair_In_split
- {var1 var1' var2 var2' T x x' y y' t a b}
- : List.In (existT _ t (a, b))
- (@flatten_binding_list
- base_type_code _ _ T
- (SmartVarfMap2 (fun t (a : var1 t) (b : var2 t) => (a, b)) x y)
- (SmartVarfMap2 (fun t (a : var1' t) (b : var2' t) => (a, b)) x' y'))
- -> List.In (existT _ t (fst a, fst b)) (@flatten_binding_list base_type_code _ _ T x x')
- /\ List.In (existT _ t (snd a, snd b)) (@flatten_binding_list base_type_code _ _ T y y').
- Proof using Type.
- induction T;
- repeat first [ exfalso; assumption
- | progress subst
- | progress inversion_sigma
- | progress inversion_prod
- | split
- | progress simpl in *
- | intro
- | progress destruct_head or
- | progress split_and
- | rewrite List.in_app_iff in *
- | solve [ eauto using List.in_or_app ] ].
- Qed.
-
- Lemma flatten_binding_list_SmartVarfMap2_pair_In_eq2_iff
- {var1 var1' var2} T x x' y y'
- : (forall t a b, List.In (existT _ t (a, b))
- (@flatten_binding_list
- base_type_code _ _ T
- (SmartVarfMap2 (fun t (a : var1 t) (b : var2 t) => (a, b)) x y)
- (SmartVarfMap2 (fun t (a : var1' t) (b : var2 t) => (a, b)) x' y'))
- -> snd a = snd b)
- <-> y = y'.
- Proof using Type.
- induction T;
- repeat first [ exfalso; assumption
- | progress subst
- | progress inversion_sigma
- | progress inversion_prod
- | progress destruct_head' unit
- | progress destruct_head' prod
- | split
- | progress simpl in *
- | intro
- | progress destruct_head or
- | apply (f_equal2 (@pair _ _))
- | progress split_iff
- | solve [ auto using List.in_or_app ]
- | match goal with
- | [ H : List.In _ (_ ++ _) |- _ ] => rewrite List.in_app_iff in H
- | [ H : context[List.In _ (_ ++ _)] |- _ ] => setoid_rewrite List.in_app_iff in H
- | [ H : forall x y, x = y -> forall t a b, List.In _ _ -> _, H' : List.In _ _ |- _ ]
- => specialize (H _ _ eq_refl _ _ _ H')
- | [ H : forall x x' y y', y = y' -> forall t a b, List.In _ _ -> _, H' : List.In _ _ |- _ ]
- => specialize (H _ _ _ _ eq_refl _ _ _ H')
- | [ H : forall t a b, _ \/ _ -> _ |- _ ]
- => pose proof (fun t a b pf => H t a b (or_introl pf));
- pose proof (fun t a b pf => H t a b (or_intror pf));
- clear H
- | [ H : forall t a b, _ |- _ ]
- => solve [ eapply (H _ (_, _) (_, _)); eauto ]
- | [ H : forall x x' y y', _ -> y = y' |- ?Y = ?Y' ]
- => specialize (fun x x' => H x x' Y Y')
- | [ H : forall x x', (forall t a b, List.In _ _ -> _ = _) -> _, H' : forall t' a' b', List.In _ _ -> _ = _ |- _ ]
- => specialize (H _ _ H')
- end ].
- Qed.
-
- Lemma flatten_binding_list_SmartVarfMap2_pair_same_in_eq2
- {var1 var1' var2} {T x x' y t a b}
- : List.In (existT _ t (a, b))
- (@flatten_binding_list
- base_type_code _ _ T
- (SmartVarfMap2 (fun t (a : var1 t) (b : var2 t) => (a, b)) x y)
- (SmartVarfMap2 (fun t (a : var1' t) (b : var2 t) => (a, b)) x' y))
- -> snd a = snd b.
- Proof using Type. intro; eapply flatten_binding_list_SmartVarfMap2_pair_In_eq2_iff; eauto. Qed.
-
- Lemma flatten_binding_list_SmartVarfMap2_pair_in_generalize2
- {var1 var1' var2 var2' var3 var3'} {T x x' y y' t a b}
- : List.In (existT _ t (a, b))
- (@flatten_binding_list
- base_type_code _ _ T
- (SmartVarfMap2 (fun t (a : var1 t) (b : var2 t) => (a, b)) x y)
- (SmartVarfMap2 (fun t (a : var1' t) (b : var2' t) => (a, b)) x' y'))
- -> (forall z z',
- exists a' b',
- List.In (existT _ t ((fst a, a'), (fst b, b')))
- (@flatten_binding_list
- base_type_code _ _ T
- (SmartVarfMap2 (fun t (a : var1 t) (b : var3 t) => (a, b)) x z)
- (SmartVarfMap2 (fun t (a : var1' t) (b : var3' t) => (a, b)) x' z'))).
- Proof.
- induction T;
- repeat first [ progress intros
- | progress subst
- | progress inversion_sigma
- | progress inversion_prod
- | progress simpl in *
- | progress destruct_head'_or
- | progress destruct_head'_prod
- | progress destruct_head'_ex
- | tauto
- | solve [ eauto ]
- | progress specialize_by_assumption
- | setoid_rewrite List.in_app_iff
- | match goal with
- | [ H : context[List.In _ (_ ++ _)] |- _ ] => setoid_rewrite List.in_app_iff in H
- | [ H : forall x : interp_flat_type ?var ?T, _, x' : interp_flat_type ?var ?T |- _ ]
- => specialize (H x')
- end ].
- Qed.
-End language.
-
-Hint Resolve wff_SmartVarf wff_SmartVarVarf wff_SmartVarVarf_nil : wf.
diff --git a/src/Compilers/WfReflective.v b/src/Compilers/WfReflective.v
deleted file mode 100644
index 4ba6d7a53..000000000
--- a/src/Compilers/WfReflective.v
+++ /dev/null
@@ -1,280 +0,0 @@
-(** * A reflective Version of [Wf] proofs *)
-(** Because every constructor of [Syntax.wff] stores the syntax tree
- being proven well-formed, a proof that a syntax tree is
- well-formed is quadratic in the size of the syntax tree. (Tacking
- an extra term on to the head of the syntax tree requires an extra
- constructor of [Syntax.wff], and that constructor stores the
- entirety of the new syntax tree.)
-
- In practice, this makes proving well-formedness of large trees
- very slow. To remedy this, we provide an alternative type
- ([reflect_wffT]) that implies [Syntax.wff], but is only linear in
- the size of the syntax tree, with a coefficient less than 1.
-
- The idea is that, since we already know the syntax-tree arguments
- to the constructors (and, moreover, already fully know the shape
- of the [Syntax.wff] proof, because it will exactly match the shape
- of the syntax tree), the proof doesn't have to store any of that
- information. It only has to store the genuinely new information
- in [Syntax.wff], namely, that the constants don't depend on the
- [var] argument (i.e., that the constants in the same location in
- the two expressions are equal), and that there are no free nor
- mismatched variables (i.e., that the variables in the same
- location in the two expressions are in the relevant list of
- binders). We can make the further optimization of storing the
- location in the list of each binder, so that all that's left to
- verify is that the locations line up correctly.
-
- Since there is no way to assign list locations (De Bruijn indices)
- after the fact (that is, once we have an [exprf var t] rather than
- an [Expr t]), we instead start from an expression where [var] is
- enriched with De Bruijn indices, and talk about [Syntax.wff] of
- that expression stripped of its De Bruijn indices. Since this
- procedure is only expected to work on concrete syntax trees, we
- will be able to simply check by unification to check that
- stripping the indices results in the term that we started with.
-
- The interface of this file is that, to prove a [Syntax.Wf] goal,
- you invoke the tactic [reflect_Wf base_type_eq_semidec_is_dec
- op_beq_bl], where:
-
- - [base_type_eq_semidec_is_dec : forall t1 t2,
- base_type_eq_semidec_transparent t1 t2 = None -> t1 <> t2] for
- some [base_type_eq_semidec_transparent : forall t1 t2 :
- base_type_code, option (t1 = t2)], and
-
- - [op_beq_bl : forall t1 tR x y, prop_of_option (op_beq t1 tR x y)
- -> x = y] for some [op_beq : forall t1 tR, op t1 tR -> op t1 tR
- -> reified_Prop] *)
-
-Require Import Coq.Arith.Arith Coq.Logic.Eqdep_dec.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.EtaWf.
-Require Import Crypto.Compilers.WfReflectiveGen.
-Require Import Crypto.Util.Notations Crypto.Util.Option Crypto.Util.Sigma Crypto.Util.Prod Crypto.Util.Decidable Crypto.Util.ListUtil.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Export Crypto.Util.PartiallyReifiedProp. (* export for the [bool >-> reified_Prop] coercion *)
-Require Export Crypto.Util.FixCoqMistakes.
-
-
-Section language.
- (** To be able to optimize away so much of the [Syntax.wff] proof,
- we must be able to decide a few things: equality of base types,
- and equality of operator codes. Since we will be casting across
- the equality proofs of base types, we require that this
- semi-decider give transparent proofs. (This requirement is not
- enforced, but it will block [vm_compute] when trying to use the
- lemma in this file.) *)
- Context (base_type_code : Type).
- Context (base_type_eq_semidec_transparent : forall t1 t2 : base_type_code, option (t1 = t2)).
- Context (base_type_eq_semidec_is_dec : forall t1 t2, base_type_eq_semidec_transparent t1 t2 = None -> t1 <> t2).
- Context (op : flat_type base_type_code -> flat_type base_type_code -> Type).
- (** In practice, semi-deciding equality of operators should either
- return [Some trivial] or [None], and not make use of the
- generality of [pointed_Prop]. However, we need to use
- [pointed_Prop] internally because we need to talk about equality
- of things of type [var t], for [var : base_type_code -> Type].
- It does not hurt to allow extra generality in [op_beq]. *)
- Context (op_beq : forall t1 tR, op t1 tR -> op t1 tR -> reified_Prop).
- Context (op_beq_bl : forall t1 tR x y, to_prop (op_beq t1 tR x y) -> x = y).
- Context {var1 var2 : base_type_code -> Type}.
-
- Local Notation eP := (fun t => var1 (fst t) * var2 (snd t))%type (only parsing).
-
- (* convenience notations that fill in some arguments used across the section *)
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
- Local Notation duplicate_type := (@duplicate_type base_type_code var1 var2).
- Local Notation reflect_wffT := (@reflect_wffT base_type_code base_type_eq_semidec_transparent op op_beq var1 var2).
- Local Notation reflect_wfT := (@reflect_wfT base_type_code base_type_eq_semidec_transparent op op_beq var1 var2).
- Local Notation flat_type_eq_semidec_transparent := (@flat_type_eq_semidec_transparent base_type_code base_type_eq_semidec_transparent).
- Local Notation preflatten_binding_list2 := (@preflatten_binding_list2 base_type_code base_type_eq_semidec_transparent var1 var2).
- Local Notation type_eq_semidec_transparent := (@type_eq_semidec_transparent base_type_code base_type_eq_semidec_transparent).
-
- Local Ltac handle_op_beq_correct :=
- repeat match goal with
- | [ H : to_prop (op_beq ?t1 ?tR ?x ?y) |- _ ]
- => apply op_beq_bl in H
- end.
- Local Ltac t_step :=
- match goal with
- | [ |- True ] => exact I
- | _ => progress cbv beta delta [eq_type_and_var op_beq' flatten_binding_list2 WfReflectiveGen.preflatten_binding_list2 option_map eq_semidec_and_gen] in *
- | _ => progress simpl in *
- | _ => progress subst
- | _ => progress break_innermost_match_step
- | _ => progress inversion_option
- | _ => progress inversion_prod
- | _ => progress inversion_reified_Prop
- | _ => congruence
- | _ => tauto
- | _ => progress intros
- | _ => progress handle_op_beq_correct
- | _ => progress specialize_by tauto
- | [ v : ex _ |- _ ] => destruct v
- | [ v : sigT _ |- _ ] => destruct v
- | [ v : prod _ _ |- _ ] => destruct v
- | [ H : forall x x', _ |- wff (flatten_binding_list ?x1 ?x2 ++ _)%list _ _ ]
- => specialize (H x1 x2)
- | [ H : forall x x', _ |- wf (existT _ _ (?x1, ?x2) :: _)%list _ _ ]
- => specialize (H x1 x2)
- | [ H : and _ _ |- _ ] => destruct H
- | [ H : to_prop (_ /\ _) |- _ ] => apply to_prop_and_reified_Prop in H; destruct H
- | [ H : context[duplicate_type (_ ++ _)%list] |- _ ]
- => rewrite duplicate_type_app in H
- | [ H : context[List.length (duplicate_type _)] |- _ ]
- => rewrite duplicate_type_length in H
- | [ H : context[List.length (_ ++ _)%list] |- _ ]
- => rewrite List.app_length in H
- | [ |- wff _ (unnatize_exprf (fst _) _) (unnatize_exprf (fst _) _) ]
- => erewrite length_natize_interp_flat_type1, length_natize_interp_flat_type2; eassumption
- | [ |- wf _ (unnatize_exprf (fst _) _) (unnatize_exprf (fst _) _) ]
- => erewrite length_natize_interp_flat_type1, length_natize_interp_flat_type2; eassumption
- | [ H : base_type_eq_semidec_transparent _ _ = None |- False ] => eapply duplicate_type_not_in; eassumption
- | [ H : List.nth_error _ _ = Some _ |- _ ] => apply List.nth_error_In in H
- | [ H : List.In _ (duplicate_type _) |- _ ] => eapply duplicate_type_in in H; [ | eassumption.. ]
- | [ H : context[match _ with _ => _ end] |- _ ] => revert H; progress break_innermost_match
- | [ |- wff _ _ _ ] => constructor
- | [ |- wf _ _ ] => constructor
- | _ => progress unfold and_reified_Prop in *
- | [ |- wff (flatten_binding_list ?x ?y) _ _ ]
- => rewrite <- (List.app_nil_r (flatten_binding_list x y))
- end.
- Local Ltac t := repeat t_step.
- Fixpoint reflect_wff (G : list (sigT (fun t => var1 t * var2 t)%type))
- {t1 t2 : flat_type}
- (e1 : @exprf (fun t => nat * var1 t)%type t1) (e2 : @exprf (fun t => nat * var2 t)%type t2)
- {struct e1}
- : let reflective_obligation := reflect_wffT (duplicate_type G) e1 e2 in
- match flat_type_eq_semidec_transparent t1 t2 with
- | Some p
- => to_prop reflective_obligation
- -> @wff base_type_code op var1 var2 G t2 (eq_rect _ exprf (unnatize_exprf (List.length G) e1) _ p) (unnatize_exprf (List.length G) e2)
- | None => True
- end.
- Proof using base_type_eq_semidec_is_dec op_beq_bl.
- cbv zeta.
- destruct e1 as [ | | ? ? ? args | tx ex tC eC | ? ex ? ey ],
- e2 as [ | | ? ? ? args' | tx' ex' tC' eC' | ? ex' ? ey' ]; simpl;
- try solve [ break_match; solve [ exact I | intros [] ] ];
- [ clear reflect_wff
- | clear reflect_wff
- | specialize (reflect_wff G _ _ args args')
- | pose proof (reflect_wff G _ _ ex ex');
- pose proof (fun x x'
- => match preflatten_binding_list2 tx tx' as v return match v with Some _ => _ | None => True end with
- | Some G0
- => reflect_wff
- (G0 x x' ++ G)%list _ _
- (eC (snd (natize_interp_flat_type (length (duplicate_type G)) x)))
- (eC' (snd (natize_interp_flat_type (length (duplicate_type G)) x')))
- | None => I
- end);
- clear reflect_wff
- | pose proof (reflect_wff G _ _ ex ex'); pose proof (reflect_wff G _ _ ey ey'); clear reflect_wff ].
- { t. }
- { t. }
- { t. }
- { t. }
- { t. }
- Qed.
- Definition reflect_wf
- {t1 t2 : type}
- (e1 : @expr (fun t => nat * var1 t)%type t1) (e2 : @expr (fun t => nat * var2 t)%type t2)
- : let reflective_obligation := reflect_wfT nil e1 e2 in
- match type_eq_semidec_transparent t1 t2 with
- | Some p
- => to_prop reflective_obligation
- -> @wf base_type_code op var1 var2 t2 (eq_rect _ expr (unnatize_expr 0 e1) _ p) (unnatize_expr 0 e2)
- | None => True
- end.
- Proof using base_type_eq_semidec_is_dec op_beq_bl.
- destruct e1 as [ tx tR f ],
- e2 as [ tx' tR' f' ]; simpl; try solve [ exact I ].
- pose proof (fun x x'
- => match preflatten_binding_list2 tx tx' as v return match v with Some _ => _ | None => True end with
- | Some G0
- => reflect_wff
- (G0 x x' ++ nil)%list
- (f (snd (natize_interp_flat_type 0 x)))
- (f' (snd (natize_interp_flat_type 0 x')))
- | None => I
- end).
- t.
- Qed.
-End language.
-
-Section Wf.
- Context (base_type_code : Type)
- (base_type_eq_semidec_transparent : forall t1 t2 : base_type_code, option (t1 = t2))
- (base_type_eq_semidec_is_dec : forall t1 t2, base_type_eq_semidec_transparent t1 t2 = None -> t1 <> t2)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type)
- (op_beq : forall t1 tR, op t1 tR -> op t1 tR -> reified_Prop)
- (op_beq_bl : forall t1 tR x y, to_prop (op_beq t1 tR x y) -> x = y)
- {t : type base_type_code}
- (e : @Expr base_type_code op t).
-
- (** Leads to smaller proofs, but is less generally applicable *)
- Theorem reflect_Wf_unnatize
- : (forall var1 var2,
- to_prop (@reflect_wfT base_type_code base_type_eq_semidec_transparent op op_beq var1 var2 nil t t (e _) (e _)))
- -> Wf (fun var => unnatize_expr 0 (e (fun t => (nat * var t)%type))).
- Proof using base_type_eq_semidec_is_dec op_beq_bl.
- intros H var1 var2; specialize (H var1 var2).
- pose proof (@reflect_wf base_type_code base_type_eq_semidec_transparent base_type_eq_semidec_is_dec op op_beq op_beq_bl var1 var2 t t (e _) (e _)) as H'.
- rewrite type_eq_semidec_transparent_refl in H' by assumption; simpl in *.
- edestruct @reflect_wfT; simpl in *; tauto.
- Qed.
-
- (** Leads to larger proofs (an extra constant factor which is the
- size of the expression tree), but more generally applicable *)
- Theorem reflect_Wf
- : (forall var1 var2,
- unnatize_expr 0 (e (fun t => (nat * var1 t)%type)) = e _
- /\ to_prop (@reflect_wfT base_type_code base_type_eq_semidec_transparent op op_beq var1 var2 nil t t (e _) (e _)))
- -> Wf e.
- Proof using base_type_eq_semidec_is_dec op_beq_bl.
- intros H var1 var2.
- rewrite <- (proj1 (H var1 var2)), <- (proj1 (H var2 var2)).
- apply reflect_Wf_unnatize, H.
- Qed.
-End Wf.
-
-(** Using [ExprEta'] ensures that reduction and conversion don't block
- on destructuring the variable arguments. *)
-Ltac preapply_eta'_Wf :=
- lazymatch goal with
- | [ |- @Wf ?base_type_code ?op ?t ?e ]
- => apply (proj1 (@Wf_ExprEta'_iff base_type_code op t e))
- end.
-Ltac generalize_reflect_Wf base_type_eq_semidec_is_dec op_beq_bl :=
- lazymatch goal with
- | [ |- @Wf ?base_type_code ?op ?t ?e ]
- => generalize (@reflect_Wf_unnatize base_type_code _ base_type_eq_semidec_is_dec op _ op_beq_bl t e)
- end.
-Ltac use_reflect_Wf :=
- let H := fresh in
- intro H;
- lazymatch type of H with
- | ?A -> ?B
- => cut A
- end;
- [ abstract vm_cast_no_check H
- | clear H ].
-Ltac fin_reflect_Wf :=
- intros;
- lazymatch goal with
- | [ |- to_prop ?P ]
- => replace P with (trueify P) by abstract vm_cast_no_check (eq_refl P)
- end;
- apply trueify_true.
-(** The tactic [reflect_Wf] is the main tactic of this file, used to
- prove [Syntax.Wf] goals *)
-Ltac reflect_Wf base_type_eq_semidec_is_dec op_beq_bl :=
- preapply_eta'_Wf;
- generalize_reflect_Wf base_type_eq_semidec_is_dec op_beq_bl;
- use_reflect_Wf; fin_reflect_Wf.
diff --git a/src/Compilers/WfReflectiveGen.v b/src/Compilers/WfReflectiveGen.v
deleted file mode 100644
index 4de378451..000000000
--- a/src/Compilers/WfReflectiveGen.v
+++ /dev/null
@@ -1,334 +0,0 @@
-(** * A reflective version of [Wf]/[WfRel] proofs *)
-(** Because every constructor of [Syntax.wff] stores the syntax tree
- being proven well-formed, a proof that a syntax tree is
- well-formed is quadratic in the size of the syntax tree. (Tacking
- an extra term on to the head of the syntax tree requires an extra
- constructor of [Syntax.wff], and that constructor stores the
- entirety of the new syntax tree.)
-
- In practice, this makes proving well-formedness of large trees
- very slow. To remedy this, we provide an alternative type
- ([reflect_wffT]) that implies [Syntax.wff], but is only linear in
- the size of the syntax tree, with a coefficient less than 1.
-
- The idea is that, since we already know the syntax-tree arguments
- to the constructors (and, moreover, already fully know the shape
- of the [Syntax.wff] proof, because it will exactly match the shape
- of the syntax tree), the proof doesn't have to store any of that
- information. It only has to store the genuinely new information
- in [Syntax.wff], namely, that the constants don't depend on the
- [var] argument (i.e., that the constants in the same location in
- the two expressions are equal), and that there are no free nor
- mismatched variables (i.e., that the variables in the same
- location in the two expressions are in the relevant list of
- binders). We can make the further optimization of storing the
- location in the list of each binder, so that all that's left to
- verify is that the locations line up correctly.
-
- Since there is no way to assign list locations (De Bruijn indices)
- after the fact (that is, once we have an [exprf var t] rather than
- an [Expr t]), we instead start from an expression where [var] is
- enriched with De Bruijn indices, and talk about [Syntax.wff] of
- that expression stripped of its De Bruijn indices. Since this
- procedure is only expected to work on concrete syntax trees, we
- will be able to simply check by unification to check that
- stripping the indices results in the term that we started with.
-
- The interface of this file is that, to prove a [Syntax.Wf] goal,
- you invoke the tactic [reflect_Wf base_type_eq_semidec_is_dec
- op_beq_bl], where:
-
- - [base_type_eq_semidec_is_dec : forall t1 t2,
- base_type_eq_semidec_transparent t1 t2 = None -> t1 <> t2] for
- some [base_type_eq_semidec_transparent : forall t1 t2 :
- base_type_code, option (t1 = t2)], and
-
- - [op_beq_bl : forall t1 tR x y, prop_of_option (op_beq t1 tR x y)
- -> x = y] for some [op_beq : forall t1 tR, op t1 tR -> op t1 tR
- -> option pointed_Prop] *)
-
-Require Import Coq.Arith.Arith Coq.Logic.Eqdep_dec.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Util.Notations Crypto.Util.Option Crypto.Util.Sigma Crypto.Util.Prod Crypto.Util.Decidable Crypto.Util.ListUtil.
-Require Import Crypto.Util.Tactics.RewriteHyp.
-Require Import Crypto.Compilers.Wf.
-Require Export Crypto.Util.PartiallyReifiedProp. (* export for the [bool >-> reified_Prop] coercion *)
-Require Export Crypto.Util.FixCoqMistakes.
-
-
-Section language.
- (** To be able to optimize away so much of the [Syntax.wff] proof,
- we must be able to decide a few things: equality of base types,
- and equality of operator codes. Since we will be casting across
- the equality proofs of base types, we require that this
- semi-decider give transparent proofs. (This requirement is not
- enforced, but it will block [vm_compute] when trying to use the
- lemma in this file.) *)
- Context (base_type_code : Type)
- (base_type_eq_semidec_transparent : forall t1 t2 : base_type_code, option (t1 = t2))
- (base_type_eq_semidec_is_dec : forall t1 t2, base_type_eq_semidec_transparent t1 t2 = None -> t1 <> t2)
- (op : flat_type base_type_code -> flat_type base_type_code -> Type).
- (** In practice, semi-deciding equality of operators should either
- return [Some trivial] or [None], and not make use of the
- generality of [pointed_Prop]. However, we need to use
- [pointed_Prop] internally because we need to talk about equality
- of things of type [var t], for [var : base_type_code -> Type].
- It does not hurt to allow extra generality in [op_beq]. *)
- Context (op_beq : forall t1 tR, op t1 tR -> op t1 tR -> reified_Prop).
- Context (op_beq_bl : forall t1 tR x y, to_prop (op_beq t1 tR x y) -> x = y).
- Context {var1 var2 : base_type_code -> Type}.
-
- Local Notation eP := (fun t => var1 (fst t) * var2 (snd t))%type (only parsing).
-
- (* convenience notations that fill in some arguments used across the section *)
- Local Notation flat_type := (flat_type base_type_code).
- Local Notation type := (type base_type_code).
- Local Notation exprf := (@exprf base_type_code op).
- Local Notation expr := (@expr base_type_code op).
-
- Local Ltac inversion_base_type_code_step :=
- match goal with
- | [ H : ?x = ?x :> base_type_code |- _ ]
- => assert (H = eq_refl) by eapply UIP_dec, dec_rel_of_semidec_rel, base_type_eq_semidec_is_dec; subst H
- | [ H : ?x = ?y :> base_type_code |- _ ] => subst x || subst y
- end.
- Local Ltac inversion_base_type_code := repeat inversion_base_type_code_step.
-
- (* lift [base_type_eq_semidec_transparent] across [flat_type] *)
- Fixpoint flat_type_eq_semidec_transparent (t1 t2 : flat_type) : option (t1 = t2)
- := match t1, t2 return option (t1 = t2) with
- | Tbase t1, Tbase t2
- => option_map (@f_equal _ _ Tbase _ _)
- (base_type_eq_semidec_transparent t1 t2)
- | Tbase _, _ => None
- | Unit, Unit => Some eq_refl
- | Unit, _ => None
- | Prod A B, Prod A' B'
- => match flat_type_eq_semidec_transparent A A', flat_type_eq_semidec_transparent B B' with
- | Some p, Some q => Some (f_equal2 Prod p q)
- | _, _ => None
- end
- | Prod _ _, _ => None
- end.
- Definition type_eq_semidec_transparent (t1 t2 : type) : option (t1 = t2)
- := match t1, t2 return option (t1 = t2) with
- | Arrow A B, Arrow A' B'
- => match flat_type_eq_semidec_transparent A A', flat_type_eq_semidec_transparent B B' with
- | Some p, Some q => Some (f_equal2 (@Arrow base_type_code) p q)
- | _, _ => None
- end
- end.
- Lemma base_type_eq_semidec_transparent_refl t : base_type_eq_semidec_transparent t t = Some eq_refl.
- Proof using base_type_eq_semidec_is_dec.
- clear -base_type_eq_semidec_is_dec.
- pose proof (base_type_eq_semidec_is_dec t t).
- destruct (base_type_eq_semidec_transparent t t); intros; try intuition congruence.
- inversion_base_type_code; reflexivity.
- Qed.
- Lemma flat_type_eq_semidec_transparent_refl t : flat_type_eq_semidec_transparent t t = Some eq_refl.
- Proof using base_type_eq_semidec_is_dec.
- clear -base_type_eq_semidec_is_dec.
- induction t as [t | | A B IHt]; simpl; try reflexivity.
- { rewrite base_type_eq_semidec_transparent_refl; reflexivity. }
- { rewrite_hyp !*; reflexivity. }
- Qed.
- Lemma type_eq_semidec_transparent_refl t : type_eq_semidec_transparent t t = Some eq_refl.
- Proof using base_type_eq_semidec_is_dec.
- clear -base_type_eq_semidec_is_dec.
- destruct t; simpl; rewrite !flat_type_eq_semidec_transparent_refl; reflexivity.
- Qed.
-
-
- Definition op_beq' t1 tR t1' tR' (x : op t1 tR) (y : op t1' tR') : reified_Prop
- := match flat_type_eq_semidec_transparent t1 t1', flat_type_eq_semidec_transparent tR tR' with
- | Some p, Some q
- => match p in (_ = t1'), q in (_ = tR') return op t1' tR' -> _ with
- | eq_refl, eq_refl => fun y => op_beq _ _ x y
- end y
- | _, _ => rFalse
- end.
-
- (** While [Syntax.wff] is parameterized over a list of [sigT (fun t
- => var1 t * var2 t)], it is simpler here to make everything
- heterogenous, rather than trying to mix homogenous and
- heterogenous things.† Thus we parameterize our [reflect_wffT]
- over a list of [sigT (fun t => var1 (fst t) * var2 (snd t))],
- and write a function ([duplicate_type]) that turns the former
- into the latter.
-
- † This is an instance of the general theme that abstraction
- barriers are important. Here we enforce the abstraction
- barrier that our input decision procedures are homogenous, but
- all of our internal code is strictly heterogenous. This
- allows us to contain the conversions between homogenous and
- heterogenous code to a few functions: [op_beq'],
- [eq_type_and_var], [eq_type_and_const], and to the statement
- about [Syntax.wff] itself. *)
-
- Definition eq_semidec_and_gen {T} (semidec : forall x y : T, option (x = y))
- (t t' : T) (f g : T -> Type) (R : forall t, f t -> g t -> reified_Prop)
- (x : f t) (x' : g t')
- : reified_Prop
- := match semidec t t' with
- | Some p
- => R _ (eq_rect _ f x _ p) x'
- | None => rFalse
- end.
-
- (* Here is where we use the generality of [pointed_Prop], to say
- that two things of type [var1] are equal, and two things of type
- [var2] are equal. *)
- Definition eq_type_and_var : sigT eP -> sigT eP -> reified_Prop
- := fun x y => (eq_semidec_and_gen
- base_type_eq_semidec_transparent _ _ var1 var1 (fun _ => rEq) (fst (projT2 x)) (fst (projT2 y))
- /\ eq_semidec_and_gen
- base_type_eq_semidec_transparent _ _ var2 var2 (fun _ => rEq) (snd (projT2 x)) (snd (projT2 y)))%reified_prop.
-
- Definition duplicate_type (ls : list (sigT (fun t => var1 t * var2 t)%type)) : list (sigT eP)
- := List.map (fun v => existT eP (projT1 v, projT1 v) (projT2 v)) ls.
-
- Lemma duplicate_type_app ls ls'
- : (duplicate_type (ls ++ ls') = duplicate_type ls ++ duplicate_type ls')%list.
- Proof using Type. apply List.map_app. Qed.
- Lemma duplicate_type_length ls
- : List.length (duplicate_type ls) = List.length ls.
- Proof using Type. apply List.map_length. Qed.
- Lemma duplicate_type_in t v ls
- : List.In (existT _ (t, t) v) (duplicate_type ls) -> List.In (existT _ t v) ls.
- Proof using base_type_eq_semidec_is_dec.
- unfold duplicate_type; rewrite List.in_map_iff.
- intros [ [? ?] [? ?] ].
- inversion_sigma; inversion_prod; inversion_base_type_code; subst; simpl.
- assumption.
- Qed.
- Lemma duplicate_type_not_in G t t0 v (H : base_type_eq_semidec_transparent t t0 = None)
- : ~List.In (existT _ (t, t0) v) (duplicate_type G).
- Proof using base_type_eq_semidec_is_dec.
- apply base_type_eq_semidec_is_dec in H.
- clear -H; intro H'.
- induction G as [|? ? IHG]; simpl in *; destruct H';
- intuition; congruence.
- Qed.
-
- Definition preflatten_binding_list2 t1 t2 : option (forall (x : interp_flat_type var1 t1) (y : interp_flat_type var2 t2), list (sigT (fun t => var1 t * var2 t)%type))
- := match flat_type_eq_semidec_transparent t1 t2 with
- | Some p
- => Some (fun x y
- => let x := eq_rect _ (interp_flat_type var1) x _ p in
- flatten_binding_list x y)
- | None => None
- end.
- Definition flatten_binding_list2 t1 t2 : option (forall (x : interp_flat_type var1 t1) (y : interp_flat_type var2 t2), list (sigT eP))
- := option_map (fun f x y => duplicate_type (f x y)) (preflatten_binding_list2 t1 t2).
- (** This function adds De Bruijn indices to variables *)
- Fixpoint natize_interp_flat_type var t (base : nat) (v : interp_flat_type var t) {struct t}
- : nat * interp_flat_type (fun t : base_type_code => nat * var t)%type t
- := match t return interp_flat_type var t -> nat * interp_flat_type _ t with
- | Prod A B => fun v => let ret := @natize_interp_flat_type _ A base (fst v) in
- let base := fst ret in
- let a := snd ret in
- let ret := @natize_interp_flat_type _ B base (snd v) in
- let base := fst ret in
- let b := snd ret in
- (base, (a, b))
- | Unit => fun v => (base, v)
- | Tbase t => fun v => (S base, (base, v))
- end v.
- Arguments natize_interp_flat_type {var t} _ _.
- Lemma length_natize_interp_flat_type1 {t} (base : nat) (v1 : interp_flat_type var1 t) (v2 : interp_flat_type var2 t)
- : fst (natize_interp_flat_type base v1) = length (flatten_binding_list v1 v2) + base.
- Proof using Type.
- revert base; induction t; simpl; [ reflexivity | reflexivity | ].
- intros; rewrite List.app_length, <- plus_assoc.
- rewrite_hyp <- ?*; reflexivity.
- Qed.
- Lemma length_natize_interp_flat_type2 {t} (base : nat) (v1 : interp_flat_type var1 t) (v2 : interp_flat_type var2 t)
- : fst (natize_interp_flat_type base v2) = length (flatten_binding_list v1 v2) + base.
- Proof using Type.
- revert base; induction t; simpl; [ reflexivity | reflexivity | ].
- intros; rewrite List.app_length, <- plus_assoc.
- rewrite_hyp <- ?*; reflexivity.
- Qed.
-
- (* This function strips De Bruijn indices from expressions *)
- Fixpoint unnatize_exprf {var t} (base : nat)
- (e : @Syntax.exprf base_type_code op (fun t => nat * var t)%type t)
- : @Syntax.exprf base_type_code op var t
- := match e in @Syntax.exprf _ _ _ t return Syntax.exprf _ _ t with
- | TT => TT
- | Var _ x => Var (snd x)
- | Op _ _ op args => Op op (@unnatize_exprf _ _ base args)
- | LetIn _ ex _ eC
- => LetIn (@unnatize_exprf _ _ base ex)
- (fun x => let v := natize_interp_flat_type base x in
- @unnatize_exprf _ _ (fst v) (eC (snd v)))
- | Pair _ x _ y
- => Pair (@unnatize_exprf _ _ base x) (@unnatize_exprf _ _ base y)
- end.
- Definition unnatize_expr {var t} (base : nat)
- (e : @Syntax.expr base_type_code op (fun t => nat * var t)%type t)
- : @Syntax.expr base_type_code op var t
- := match e in @Syntax.expr _ _ _ t return Syntax.expr _ _ t with
- | Abs tx tR f => Abs (fun x : interp_flat_type var tx =>
- let v := natize_interp_flat_type (t:=tx) base x in
- @unnatize_exprf _ _ (fst v) (f (snd v)))
- end.
-
- Fixpoint reflect_wffT (G : list (sigT (fun t => var1 (fst t) * var2 (snd t))%type))
- {t1 t2 : flat_type}
- (e1 : @exprf (fun t => nat * var1 t)%type t1)
- (e2 : @exprf (fun t => nat * var2 t)%type t2)
- {struct e1}
- : reified_Prop
- := match e1, e2 with
- | TT, TT => rTrue
- | TT, _ => rFalse
- | Var t0 x, Var t1 y
- => match beq_nat (fst x) (fst y), List.nth_error G (List.length G - S (fst x)) with
- | true, Some v => eq_type_and_var v (existT _ (t0, t1) (snd x, snd y))
- | _, _ => rFalse
- end
- | Var _ _, _ => rFalse
- | Op t1 tR op args, Op t1' tR' op' args'
- => (@reflect_wffT G t1 t1' args args' /\ op_beq' t1 tR t1' tR' op op')%reified_prop
- | Op _ _ _ _, _ => rFalse
- | LetIn tx ex tC eC, LetIn tx' ex' tC' eC'
- => let p := @reflect_wffT G tx tx' ex ex' in
- match @flatten_binding_list2 tx tx', flat_type_eq_semidec_transparent tC tC' with
- | Some G0, Some _
- => p
- /\ (∀ (x : interp_flat_type var1 tx) (x' : interp_flat_type var2 tx'),
- @reflect_wffT (G0 x x' ++ G)%list _ _
- (eC (snd (natize_interp_flat_type (List.length G) x)))
- (eC' (snd (natize_interp_flat_type (List.length G) x'))))
- | _, _ => rFalse
- end
- | LetIn _ _ _ _, _ => rFalse
- | Pair tx ex ty ey, Pair tx' ex' ty' ey'
- => @reflect_wffT G tx tx' ex ex' /\ @reflect_wffT G ty ty' ey ey'
- | Pair _ _ _ _, _ => rFalse
- end%reified_prop.
-
- Definition reflect_wfT (G : list (sigT (fun t => var1 (fst t) * var2 (snd t))%type))
- {t1 t2 : type}
- (e1 : @expr (fun t => nat * var1 t)%type t1)
- (e2 : @expr (fun t => nat * var2 t)%type t2)
- : reified_Prop
- := match e1, e2 with
- | Abs tx tR f, Abs tx' tR' f'
- => match @flatten_binding_list2 tx tx', flat_type_eq_semidec_transparent tR tR' with
- | Some G0, Some _
- => ∀ (x : interp_flat_type var1 tx) (x' : interp_flat_type var2 tx'),
- @reflect_wffT (G0 x x' ++ G)%list _ _
- (f (snd (natize_interp_flat_type (List.length G) x)))
- (f' (snd (natize_interp_flat_type (List.length G) x')))
- | _, _ => rFalse
- end
- end%reified_prop.
-End language.
-
-Global Arguments reflect_wffT {_} _ {op} op_beq {var1 var2} G {t1 t2} _ _.
-Global Arguments reflect_wfT {_} _ {op} op_beq {var1 var2} G {t1 t2} _ _.
-Global Arguments unnatize_exprf {_ _ _ _} _ _.
-Global Arguments unnatize_expr {_ _ _ _} _ _.
-Global Arguments natize_interp_flat_type {_ _ t} _ _.
diff --git a/src/Compilers/Z/ArithmeticSimplifier.v b/src/Compilers/Z/ArithmeticSimplifier.v
deleted file mode 100644
index 5a5b95a22..000000000
--- a/src/Compilers/Z/ArithmeticSimplifier.v
+++ /dev/null
@@ -1,611 +0,0 @@
-(** * SimplifyArith: Remove things like (_ * 1), (_ + 0), etc *)
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Rewriter.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Equality.
-Require Import Crypto.Util.ZUtil.Definitions.
-
-Section language.
- Context (convert_adc_to_sbb : bool).
- Local Notation exprf := (@exprf 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
- -> match dst with
- | Tbase t => inverted_expr t
- | Prod _ _ => True
- | _ => inverted_expr TZ
- end
- with
- | OpConst _ z => fun _ _ => const_of _ z
- | Opp TZ TZ => fun _ args => neg_expr _ args
- | MulSplit _ _ _ _ _ => fun _ _ => I
- | AddWithGetCarry _ _ _ _ _ _ => fun _ _ => I
- | SubWithGetBorrow _ _ _ _ _ _ => fun _ _ => I
- | _ => 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.
-
- Fixpoint uninterp_expr_or_const {t} : interp_flat_type inverted_expr t -> exprf (var:=var) t
- := match t with
- | Tbase T => fun e => match e with
- | const_of v => Op (OpConst v) TT
- | gen_expr e => e
- | neg_expr e => Op (Opp _ _) e
- end
- | Unit => fun _ => TT
- | Prod A B => fun (ab : interp_flat_type _ A * interp_flat_type _ B)
- => Pair (@uninterp_expr_or_const A (fst ab))
- (@uninterp_expr_or_const B (snd ab))
- 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
- | Add T1 T2 Tout as opc
- => fun args
- => match interp_as_expr_or_const args with
- | Some (const_of v, gen_expr e)
- => if (v =? 0)%Z
- then match base_type_eq_semidec_transparent T2 Tout with
- | Some pf => eq_rect _ (fun t => exprf (Tbase t)) e _ pf
- | None => Op opc args
- end
- else Op opc args
- | Some (gen_expr e, const_of v)
- => if (v =? 0)%Z
- then match base_type_eq_semidec_transparent T1 Tout with
- | Some pf => eq_rect _ (fun t => exprf (Tbase t)) e _ pf
- | None => Op opc args
- end
- else Op opc args
- | _ => 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 if (v >? 0)%Z
- then Op (Opp _ _) (Op opc (Pair (Op (OpConst v) TT) 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
- | Mul (TWord bw1 as T1) (TWord bw2 as T2) (TWord bwout as Tout) as opc
- => fun args
- => let sz1 := (2^Z.of_nat (2^bw1))%Z in
- let sz2 := (2^Z.of_nat (2^bw2))%Z in
- let szout := (2^Z.of_nat (2^bwout))%Z in
- match interp_as_expr_or_const args with
- | Some (const_of l, const_of r)
- => Op (OpConst (((Z.max 0 l mod sz1) * (Z.max 0 r mod sz2)) mod szout)%Z) TT
- | Some (const_of v, gen_expr e)
- => if ((Z.max 0 v mod sz1) mod szout =? 0)%Z
- then Op (OpConst 0%Z) TT
- else if ((Z.max 0 v mod sz1) mod szout =? 1)%Z
- then match base_type_eq_semidec_transparent T2 Tout with
- | Some pf => eq_rect _ (fun t => exprf (Tbase t)) e _ pf
- | None => Op opc args
- end
- else Op opc args
- | Some (gen_expr e, const_of v)
- => if ((Z.max 0 v mod sz2) mod szout =? 0)%Z
- then Op (OpConst 0%Z) TT
- else if ((Z.max 0 v mod sz2) mod szout =? 1)%Z
- then match base_type_eq_semidec_transparent T1 Tout with
- | Some pf => eq_rect _ (fun t => exprf (Tbase t)) e _ pf
- | None => Op opc args
- end
- else Op opc args
- | _ => 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
- | MulSplit bitwidth (TWord bw1 as T1) (TWord bw2 as T2) (TWord bwout1 as Tout1) (TWord bwout2 as Tout2) as opc
- => fun args
- => let sz1 := (2^Z.of_nat (2^bw1))%Z in
- let sz2 := (2^Z.of_nat (2^bw2))%Z in
- let szout1 := (2^Z.of_nat (2^bwout1))%Z in
- let szout2 := (2^Z.of_nat (2^bwout2))%Z in
- match interp_as_expr_or_const args with
- | Some (const_of l, const_of r)
- => let '(a, b) := Z.mul_split_at_bitwidth bitwidth (Z.max 0 l mod sz1) (Z.max 0 r mod sz2) in
- Pair (Op (OpConst (a mod szout1)%Z) TT)
- (Op (OpConst (b mod szout2)%Z) TT)
- | Some (const_of v, gen_expr e)
- => let v' := (Z.max 0 v mod sz1)%Z in
- if (v' =? 0)%Z
- then Pair (Op (OpConst 0%Z) TT) (Op (OpConst 0%Z) TT)
- else if ((v' =? 1) && (2^Z.of_nat (2^bw2) <=? 2^bitwidth))%Z%bool
- then match base_type_eq_semidec_transparent T2 Tout1 with
- | Some pf => Pair (eq_rect _ (fun t => exprf (Tbase t)) e _ pf)
- (Op (OpConst 0%Z) TT)
- | None => Op opc args
- end
- else Op opc args
- | Some (gen_expr e, const_of v)
- => let v' := (Z.max 0 v mod sz2)%Z in
- if (v' =? 0)%Z
- then Pair (Op (OpConst 0%Z) TT) (Op (OpConst 0%Z) TT)
- else if ((v' =? 1) && (2^Z.of_nat (2^bw1) <=? 2^bitwidth))%Z%bool
- then match base_type_eq_semidec_transparent T1 Tout1 with
- | Some pf => Pair (eq_rect _ (fun t => exprf (Tbase t)) e _ pf)
- (Op (OpConst 0%Z) TT)
- | None => Op opc args
- end
- else Op opc args
- | _ => Op opc args
- end
- | IdWithAlt (TWord _ as T1) _ (TWord _ as Tout) as opc
- => fun args
- => match base_type_eq_semidec_transparent T1 Tout with
- | Some pf
- => match interp_as_expr_or_const args with
- | Some (const_of c, _)
- => Op (OpConst c) TT
- | Some (neg_expr e, _)
- => Op (Opp _ _) e
- | Some (gen_expr e, _)
- => eq_rect _ (fun t => exprf (Tbase t)) e _ pf
- | None
- => Op opc args
- end
- | None
- => Op opc args
- end
- | IdWithAlt TZ TZ TZ as opc
- => fun args
- => match interp_as_expr_or_const args with
- | Some (gen_expr e1, gen_expr e2)
- => match invert_Op e1, invert_Op e2 with
- | Some (existT _ (Add TZ TZ TZ as opc1, args1)),
- Some (existT _ (Add TZ TZ TZ as opc2, args2))
- | Some (existT _ (Sub TZ TZ TZ as opc1, args1)),
- Some (existT _ (Sub TZ TZ TZ as opc2, args2))
- | Some (existT _ (Mul TZ TZ TZ as opc1, args1)),
- Some (existT _ (Mul TZ TZ TZ as opc2, args2))
- => match interp_as_expr_or_const args1, interp_as_expr_or_const args2 with
- | Some (gen_expr e1, const_of c1),
- Some (gen_expr e2, const_of c2)
- => if Z.eqb c1 c2
- then Op opc1 (Op opc (e1, e2), Op (OpConst c1) TT)%expr
- else Op opc args
- | _, _
- => Op opc args
- end
- | _, _
- => Op opc args
- end
- | Some (neg_expr e1, neg_expr e2)
- => Op (Opp _ _) (Op opc (e1, e2)%expr)
- | Some (const_of c1, const_of c2)
- => Op (OpConst c1) TT
- | _
- => Op opc args
- end
- | IdWithAlt _ _ (TWord _ as Tout) as opc
- => fun args
- => match interp_as_expr_or_const args with
- | Some (gen_expr e1, _)
- => match invert_Op e1 with
- | Some (existT _ (Add (TWord _) (TWord _) TZ as opc1, args1))
- => Op (Add _ _ Tout) args1
- | Some (existT _ (Sub (TWord _) (TWord _) TZ as opc1, args1))
- => Op (Sub _ _ Tout) args1
- | Some (existT _ (Mul (TWord _) (TWord _) TZ as opc1, args1))
- => Op (Mul _ _ Tout) args1
- | _
- => Op opc args
- end
- | _
- => Op opc args
- end
- | Zselect TZ TZ TZ TZ as opc
- => fun args
- => match interp_as_expr_or_const args with
- | Some (const_of c, x, y)
- => match (c =? 0)%Z, x, y with
- | true, const_of c, _
- | false, _, const_of c
- => Op (OpConst c) TT
- | true, gen_expr e, _
- | false, _, gen_expr e
- => e
- | true, neg_expr e, _
- | false, _, neg_expr e
- => Op (Opp TZ TZ) e
- end
- | Some (neg_expr e, x, y)
- => let x := uninterp_expr_or_const (t:=Tbase _) x in
- let y := uninterp_expr_or_const (t:=Tbase _) y in
- Op (Zselect TZ TZ TZ TZ) (e, x, y)%expr
- | _ => Op opc args
- end
- | AddWithCarry TZ TZ TZ TZ as opc
- => fun args
- => let first_pass
- := match interp_as_expr_or_const args with
- | Some (const_of c, const_of x, const_of y)
- => Some (Op (OpConst (interp_op _ _ opc (c, x, y))) TT)
- | Some (gen_expr e, const_of c1, const_of c2)
- | Some (const_of c1, gen_expr e, const_of c2)
- | Some (const_of c1, const_of c2, gen_expr e)
- => if (c1 + c2 =? 0)%Z
- then Some e
- else None
- | _ => None
- end in
- match first_pass with
- | Some e => e
- | None
- => if convert_adc_to_sbb
- then match interp_as_expr_or_const args with
- | Some (const_of c, const_of x, const_of y)
- => Op (OpConst (interp_op _ _ opc (c, x, y))) TT
- | Some (c, gen_expr x, y)
- => let y' := match y with
- | const_of y => if (y <? 0)%Z
- then Some (Op (OpConst (-y)) TT)
- else None
- | neg_expr y => Some y
- | gen_expr _ => None
- end in
- match y' with
- | Some y => Op (SubWithBorrow TZ TZ TZ TZ)
- (match c with
- | const_of c => Op (OpConst (-c)) TT
- | neg_expr c => c
- | gen_expr c => Op (Opp TZ TZ) c
- end,
- x, y)%expr
- | None => Op opc args
- end
- | _ => Op opc args
- end
- else Op opc args
- end
- | AddWithGetCarry bw TZ TZ TZ TZ TZ as opc
- => fun args
- => if convert_adc_to_sbb
- then match interp_as_expr_or_const args with
- | Some (const_of c, const_of x, const_of y)
- => let '(v, c) := interp_op _ _ opc (c, x, y) in
- (Op (OpConst v) TT, Op (OpConst c) TT)%expr
- | Some (c, gen_expr x, y)
- => let y' := match y with
- | const_of y => if (y <? 0)%Z
- then Some (Op (OpConst (-y)) TT)
- else None
- | neg_expr y => Some y
- | gen_expr _ => None
- end in
- let c' := match c with
- | const_of c => if (c <? 0)%Z
- then Some (Op (OpConst (-c)) TT)
- else None
- | neg_expr c => Some c
- | gen_expr _ => None
- end in
- match c', y' with
- | _, Some y => LetIn (Op (SubWithGetBorrow bw TZ TZ TZ TZ TZ)
- (match c with
- | const_of c => Op (OpConst (-c)) TT
- | neg_expr c => c
- | gen_expr c => Op (Opp TZ TZ) c
- end,
- x, y)%expr)
- (fun '(v, c) => (Var v, Op (Opp TZ TZ) (Var c))%expr)
- | Some c, _ => LetIn (Op (SubWithGetBorrow bw TZ TZ TZ TZ TZ)
- (c,
- x,
- match y with
- | const_of y => Op (OpConst (-y)) TT
- | neg_expr y => y
- | gen_expr y => Op (Opp TZ TZ) y
- end)%expr)
- (fun '(v, c) => (Var v, Op (Opp TZ TZ) (Var c))%expr)
- | None, None => Op opc args
- end
- | Some (c, const_of x, y)
- => let y' := match y with
- | const_of y => if (y <? 0)%Z
- then Some (Op (OpConst (-y)) TT)
- else None
- | neg_expr y => Some y
- | gen_expr _ => None
- end in
- let c' := match c with
- | const_of c => if (c <? 0)%Z
- then Some (Op (OpConst (-c)) TT)
- else None
- | neg_expr c => Some c
- | gen_expr _ => None
- end in
- match c', y' with
- | _, Some y => LetIn (Op (SubWithGetBorrow bw TZ TZ TZ TZ TZ)
- (match c with
- | const_of c => Op (OpConst (-c)) TT
- | neg_expr c => c
- | gen_expr c => Op (Opp TZ TZ) c
- end,
- Op (OpConst x) TT, y)%expr)
- (fun '(v, c) => (Var v, Op (Opp TZ TZ) (Var c))%expr)
- | Some c, _ => LetIn (Op (SubWithGetBorrow bw TZ TZ TZ TZ TZ)
- (c,
- Op (OpConst x) TT,
- match y with
- | const_of y => Op (OpConst (-y)) TT
- | neg_expr y => y
- | gen_expr y => Op (Opp TZ TZ) y
- end)%expr)
- (fun '(v, c) => (Var v, Op (Opp TZ TZ) (Var c))%expr)
- | None, None => Op opc args
- end
- | _ => Op opc args
- end
- else Op opc args
- | AddWithGetCarry bw (TWord bw1 as T1) (TWord bw2 as T2) (TWord bw3 as T3) (TWord bwout as Tout) Tout2 as opc
- => fun args
- => let pass0
- := if ((0 <=? bw)%Z && (2^Z.of_nat (2^bw1) + 2^Z.of_nat (2^bw2) + 2^Z.of_nat (2^bw3) - 3 <=? 2^bw - 1)%Z)%nat%bool
- then Some (Pair (LetIn args (fun '(a, b, c) => Op (Add _ _ _) (Pair (Op (Add _ _ Tout) (Pair (Var a) (Var b))) (Var c))))
- (Op (OpConst 0) TT))
- else None
- in
- match pass0 with
- | Some e => e
- | None
- => match interp_as_expr_or_const args with
- | Some (const_of c, const_of x, const_of y)
- => if ((c =? 0) && (x =? 0) && (y =? 0))%Z%bool
- then Pair (Op (OpConst 0) TT) (Op (OpConst 0) TT)
- else Op opc args
- | Some (gen_expr e, const_of c1, const_of c2)
- => match base_type_eq_semidec_transparent T1 Tout with
- | Some pf
- => if ((c1 =? 0) && (c2 =? 0) && (2^Z.of_nat bw1 <=? bw))%Z%bool
- then Pair (eq_rect _ (fun t => exprf (Tbase t)) e _ pf) (Op (OpConst 0) TT)
- else Op opc args
- | None
- => Op opc args
- end
- | Some (const_of c1, gen_expr e, const_of c2)
- => match base_type_eq_semidec_transparent T2 Tout with
- | Some pf
- => if ((c1 =? 0) && (c2 =? 0) && (2^Z.of_nat bw2 <=? bw))%Z%bool
- then Pair (eq_rect _ (fun t => exprf (Tbase t)) e _ pf) (Op (OpConst 0) TT)
- else Op opc args
- | None
- => Op opc args
- end
- | Some (const_of c1, const_of c2, gen_expr e)
- => match base_type_eq_semidec_transparent T3 Tout with
- | Some pf
- => if ((c1 =? 0) && (c2 =? 0) && (2^Z.of_nat bw3 <=? bw))%Z%bool
- then Pair (eq_rect _ (fun t => exprf (Tbase t)) e _ pf) (Op (OpConst 0) TT)
- else Op opc args
- | None
- => Op opc args
- end
- | _ => Op opc args
- end
- end
- | SubWithBorrow TZ TZ TZ TZ as opc
- => fun args
- => if convert_adc_to_sbb
- then match interp_as_expr_or_const args with
- | Some (const_of c, const_of x, const_of y)
- => Op (OpConst (interp_op _ _ opc (c, x, y))) TT
- | Some (c, gen_expr x, y)
- => let y' := match y with
- | const_of y => if (y <? 0)%Z
- then Some (Op (OpConst (-y)) TT)
- else None
- | neg_expr y => Some y
- | gen_expr _ => None
- end in
- match y' with
- | Some y => Op (AddWithCarry TZ TZ TZ TZ)
- (match c with
- | const_of c => Op (OpConst (-c)) TT
- | neg_expr c => c
- | gen_expr c => Op (Opp TZ TZ) c
- end,
- x, y)%expr
- | None => Op opc args
- end
- | _ => Op opc args
- end
- else Op opc args
- | SubWithGetBorrow bw TZ TZ TZ TZ TZ as opc
- => fun args
- => if convert_adc_to_sbb
- then match interp_as_expr_or_const args with
- | Some (const_of c, const_of x, const_of y)
- => let '(v, c) := interp_op _ _ opc (c, x, y) in
- (Op (OpConst v) TT, Op (OpConst c) TT)%expr
- | Some (c, gen_expr x, y)
- => let y' := match y with
- | const_of y => if (y <? 0)%Z
- then Some (Op (OpConst (-y)) TT)
- else None
- | neg_expr y => Some y
- | gen_expr _ => None
- end in
- match y' with
- | Some y => LetIn (Op (AddWithGetCarry bw TZ TZ TZ TZ TZ)
- (match c with
- | const_of c => Op (OpConst (-c)) TT
- | neg_expr c => c
- | gen_expr c => Op (Opp TZ TZ) c
- end,
- x, y)%expr)
- (fun '(v, c) => (Var v, Op (Opp TZ TZ) (Var c))%expr)
- | None => Op opc args
- end
- | _ => Op opc args
- end
- else Op opc args
- | Sub _ _ _ as opc
- | Mul _ _ _ as opc
- | Shl _ _ _ as opc
- | Shr _ _ _ as opc
- | Land _ _ _ as opc
- | Lor _ _ _ as opc
- | OpConst _ _ as opc
- | Opp _ _ as opc
- | IdWithAlt _ _ _ as opc
- | Zselect _ _ _ _ as opc
- | MulSplit _ _ _ _ _ as opc
- | AddWithCarry _ _ _ _ as opc
- | AddWithGetCarry _ _ _ _ _ _ as opc
- | SubWithBorrow _ _ _ _ as opc
- | SubWithGetBorrow _ _ _ _ _ _ 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
deleted file mode 100644
index c1c841c9f..000000000
--- a/src/Compilers/Z/ArithmeticSimplifierInterp.v
+++ /dev/null
@@ -1,247 +0,0 @@
-Require Import Coq.micromega.Psatz.
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-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.Hints.
-Require Import Crypto.Util.ZUtil.Hints.Core.
-Require Import Crypto.Util.ZUtil.ZSimplify.Core.
-Require Import Crypto.Util.ZUtil.Definitions.
-Require Import Crypto.Util.ZUtil.Z2Nat.
-Require Import Crypto.Util.ZUtil.AddGetCarry.
-Require Import Crypto.Util.ZUtil.Modulo.PullPush.
-Require Import Crypto.Util.ZUtil.Tactics.LtbToLt.
-Require Import Crypto.Util.ZUtil.Tactics.ZeroBounds.
-Require Import Crypto.Util.ZUtil.Tactics.RewriteModSmall.
-Require Import Crypto.Util.ZUtil.Tactics.PullPush.Modulo.
-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.
-Require Import Crypto.Util.Tactics.UniquePose.
-
-Local Notation exprf := (@exprf base_type op interp_base_type).
-Local Notation expr := (@expr base_type op interp_base_type).
-
-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 subst_prod
- | progress destruct_head'_and
- | progress destruct_head'_prod
- | progress eliminate_hprop_eq
- | progress break_innermost_match_step
- | progress break_match_hyps ].
-
-Local Ltac solve_word_small _ :=
- lazymatch goal with
- | [ H : (2^Z.of_nat ?b <= ?bw)%Z |- (0 <= FixedWordSizes.wordToZ ?x < 2^?bw)%Z ]
- => cut (0 <= FixedWordSizes.wordToZ x < 2^(Z.of_nat (2^b)%nat))%Z;
- [ rewrite Z.pow_Zpow; cbn [Z.of_nat Pos.of_succ_nat Pos.succ];
- assert ((2^2^Z.of_nat b <= 2^bw)%Z) by auto with zarith;
- auto with zarith
- | apply FixedWordSizesEquality.wordToZ_range ]
- end.
-
-Definition interpf_as_expr_or_const {t}
- : interp_flat_type (@inverted_expr interp_base_type) t -> interp_flat_type interp_base_type t
- := SmartVarfMap
- (fun t z => 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).
-
-Lemma interp_as_expr_or_const_correct {t} e z
- : @interp_as_expr_or_const interp_base_type t e = Some z
- -> interpf interp_op e = interpf_as_expr_or_const z.
-Proof.
- induction e;
- repeat first [ progress subst
- | progress inversion_option
- | progress simpl in *
- | progress cbn [interpf_as_expr_or_const SmartVarfMap smart_interp_flat_map]
- | reflexivity
- | break_innermost_match_hyps_step
- | intro
- | match goal with
- | [ H : forall z, Some _ = Some z -> _ |- _ ] => specialize (H _ eq_refl)
- | [ H : interpf _ ?e = interpf_as_expr_or_const _ |- _ ]
- => rewrite H
- | [ |- context[match ?e with _ => _ end] ]
- => is_var e; invert_one_op e
- end
- | break_innermost_match_step ].
-Qed.
-
-Local Ltac rewrite_interp_as_expr_or_const_correct _ :=
- match goal with
- | [ |- context[interpf _ ?e] ]
- => erewrite !(@interp_as_expr_or_const_correct _ e) by eassumption; cbv beta iota;
- cbn [interpf_as_expr_or_const SmartVarfMap smart_interp_flat_map]
- end.
-
-Local Arguments Z.mul !_ !_.
-Local Arguments Z.add !_ !_.
-Local Arguments Z.sub !_ !_.
-Local Arguments Z.opp !_.
-Local Arguments interp_op _ _ !_ _ / .
-Local Arguments lift_op / .
-Local Opaque Z.pow.
-
-Lemma InterpSimplifyArith {convert_adc_to_sbb} {t} (e : Expr t)
- : forall x, Interp (SimplifyArith convert_adc_to_sbb e) x = Interp e x.
-Proof.
- apply InterpRewriteOp; intros; unfold simplify_op_expr.
- Time break_innermost_match;
- repeat first [ reflexivity
- | progress subst
- | progress simpl in *
- | progress inversion_prod
- | progress invert_expr_subst
- | inversion_base_type_constr_step
- | match goal with
- | [ |- context[match ?e with _ => _ end] ]
- => is_var e; invert_one_op e;
- repeat match goal with
- | [ |- match ?T with _ => _ end _ ]
- => break_innermost_match_step; try exact I
- end
- end
- | break_innermost_match_step
- | rewrite_interp_as_expr_or_const_correct ()
- | intro ].
- all:repeat first [ reflexivity
- | omega
- | discriminate
- | progress cbv [LetIn.Let_In Z.zselect IdfunWithAlt.id_with_alt]
- | progress subst
- | progress simpl in *
- | progress Bool.split_andb
- | progress Z.ltb_to_lt
- | break_innermost_match_step
- | apply (f_equal2 pair)
- | progress cbv [cast_const ZToInterp interpToZ]
- | match goal with
- | [ |- interpf ?interp_op ?e = ?x ]
- => rewrite <- (FixedWordSizesEquality.ZToWord_wordToZ (interpf interp_op e)), <- FixedWordSizesEquality.eq_ZToWord
- end
- | rewrite <- FixedWordSizesEquality.eq_ZToWord ].
- all:repeat first [ rewrite FixedWordSizesEquality.ZToWord_wordToZ
- | rewrite FixedWordSizesEquality.ZToWord_wordToZ_ZToWord by reflexivity
- | rewrite FixedWordSizesEquality.wordToZ_ZToWord_0
- | rewrite !FixedWordSizesEquality.wordToZ_ZToWord_mod_full ].
- all:repeat match goal with
- | [ H : _ = Some eq_refl |- _ ] => clear H
- | [ H : interp_as_expr_or_const _ = Some _ |- _ ] => clear H
- | [ H : interpf _ _ = _ |- _ ] => clear H
- | [ H : Syntax.exprf _ _ _ |- _ ] => clear H
- | [ H : Expr _ |- _ ] => clear H
- | [ H : type _ |- _ ] => clear H
- | [ H : bool |- _ ] => clear H
- | [ |- context[FixedWordSizes.wordToZ ?e] ]
- => pose proof (FixedWordSizesEquality.wordToZ_range e);
- lazymatch e with
- | interpf interp_op ?e'
- => generalize dependent (FixedWordSizes.wordToZ e); clear e'; intros
- | _ => is_var e; generalize dependent (FixedWordSizes.wordToZ e);
- clear e; intros
- end
- | [ |- context[interpf interp_op ?e] ]
- => is_var e; generalize dependent (interpf interp_op e); clear e; intros
- | [ |- context[Z.of_nat (2^?e)] ]
- => is_var e; assert ((0 < Z.of_nat (2^e))%Z)
- by (rewrite Z.pow_Zpow; simpl Z.of_nat; Z.zero_bounds);
- generalize dependent (Z.of_nat (2^e)); clear e; intros
- end.
- all:try nia.
- Time
- all:repeat first [ reflexivity
- | omega
- | progress change (2^0)%Z with 1%Z in *
- | progress change (2^1)%Z with 2%Z in *
- | progress rewrite ?Z.land_0_l, ?Z.land_0_r, ?Z.lor_0_l, ?Z.lor_0_r, ?Z.opp_involutive, ?Z.shiftr_0_r
- | progress rewrite ?Z.land_ones by lia
- | match goal with
- | [ H : (?bitwidth < 0)%Z, H' : context[(2^?bitwidth)%Z] |- _ ] => is_var bitwidth; destruct bitwidth; cbn in H
- end
- | progress autorewrite with Zshift_to_pow in *
- | rewrite !Z.sub_with_borrow_to_add_get_carry
- | progress cbv [Z.add_with_carry]
- | rewrite Z.mod_mod by Z.zero_bounds
- | match goal with
- | [ |- context[(?x mod ?y)%Z] ]
- => lazymatch goal with
- | [ H : (0 <= x mod y)%Z |- _ ] => fail
- | [ H : (0 <= x mod y < _)%Z |- _ ] => fail
- | _ => assert (0 <= x mod y < y)%Z by (apply Z.mod_pos_bound; Z.zero_bounds; lia)
- end
- | [ |- context[(?x / ?y)%Z] ]
- => lazymatch goal with
- | [ H : (0 <= x / y)%Z |- _ ] => fail
- | _ => assert (0 <= x / y)%Z by Z.zero_bounds
- end
- | [ H : (2^Z.of_nat ?bw <= ?bw')%Z |- context[(2^?bw')%Z] ]
- => unique assert ((2^Z.of_nat (2^bw) <= 2^bw')%Z)
- by (rewrite Z.pow_Zpow; simpl @Z.of_nat; auto with zarith)
- end
- | progress autorewrite with zsimplify_const in *
- | match goal with
- | [ H : (2^?x <= 1)%Z, H' : (0 < ?x)%Z |- _ ]
- => lazymatch goal with
- | [ |- False ] => fail
- | _ => exfalso; clear -H H'; assert (2^1 <= 2^x)%Z by auto with zarith
- end
- | [ H : (0 <= ?x < _)%Z |- context[Z.max 0 ?x] ]
- => rewrite (Z.max_r 0 x) in * by apply H
- | [ H : (0 <= ?x)%Z |- context[Z.max 0 ?x] ]
- => rewrite (Z.max_r 0 x) in * by apply H
- | [ H : (0 <= ?x < _)%Z, H' : (0 <= ?y < _)%Z |- context[Z.max 0 (?x * ?y)] ]
- => rewrite (Z.max_r 0 (x * y)) in * by (apply Z.mul_nonneg_nonneg; first [ apply H | apply H' ])
- | [ H : (0 <= ?x < _)%Z, H' : (0 <= ?y < _)%Z |- context[Z.max 0 (?x + ?y)] ]
- => rewrite (Z.max_r 0 (x + y)) in * by (apply Z.add_nonneg_nonneg; first [ apply H | apply H' ])
- | [ H : ?x = 0%Z |- context[?x] ] => rewrite H
- | [ H : ?x = 0%Z, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : ?x = Z.pos _ |- context[?x] ] => rewrite H
- | [ H : ?x = Z.pos _, H' : context[?x] |- _ ] => rewrite H in H'
- | [ H : context[(_^Z.neg ?p)%Z] |- _ ]
- => rewrite (Z.pow_neg_r _ (Z.neg p)) in H by lia
- | [ H : (?x mod ?y = 0)%Z |- context[((?x * _) mod ?y)%Z] ]
- => rewrite (Z.mul_mod_full x _ y)
- | [ H : (?x mod ?y = 0)%Z |- context[((_ * ?x) mod ?y)%Z] ]
- => rewrite (Z.mul_mod_full _ x y)
- | [ H : ?x = Z.pos _ |- context[?x] ] => rewrite H
- | [ H : (?x mod ?y = Z.pos _)%Z |- context[((?x * _) mod ?y)%Z] ]
- => rewrite (Z.mul_mod_full x _ y)
- | [ H : (?x mod ?y = Z.pos _)%Z |- context[((_ * ?x) mod ?y)%Z] ]
- => rewrite (Z.mul_mod_full _ x y)
- | [ |- context[(?x mod ?m)%Z] ]
- => rewrite (Z.mod_small x m) by Z.rewrite_mod_small_solver
- | [ |- context[(?x / ?m)%Z] ]
- => rewrite (Z.div_small x m) by Z.rewrite_mod_small_solver
- end
- | progress pull_Zmod ].
-Qed.
-
-Hint Rewrite @InterpSimplifyArith : reflective_interp.
diff --git a/src/Compilers/Z/ArithmeticSimplifierUtil.v b/src/Compilers/Z/ArithmeticSimplifierUtil.v
deleted file mode 100644
index 49d3a2257..000000000
--- a/src/Compilers/Z/ArithmeticSimplifierUtil.v
+++ /dev/null
@@ -1,79 +0,0 @@
-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
deleted file mode 100644
index 4efa7445a..000000000
--- a/src/Compilers/Z/ArithmeticSimplifierWf.v
+++ /dev/null
@@ -1,218 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-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.Compilers.Z.Syntax.Util.
-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 wff := (@wff base_type op).
-Local Notation Wf := (@Wf base_type op).
-
-Local Ltac fin_t :=
- first [ exact I
- | reflexivity
- | congruence
- | assumption
- | exfalso; assumption
- | match goal with
- | [ |- _ /\ False ] => exfalso
- | [ |- False /\ _ ] => exfalso
- | [ |- _ /\ _ /\ False ] => exfalso
- | [ |- _ /\ False /\ _ ] => exfalso
- | [ |- False /\ _ /\ _ ] => exfalso
- end ].
-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_hyps_step
- | 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.
-
-Local Ltac pret_step :=
- first [ fin_t
- | progress subst
- | progress inversion_option
- | progress inversion_prod
- | progress simpl in *
- | progress inversion_wf
- | match goal with
- | [ H : match interp_as_expr_or_const ?e with _ => _ end = Some _ |- _ ]
- => is_var e; destruct (interp_as_expr_or_const e) eqn:?
- end ].
-
-Fixpoint wff_as_expr_or_const {var1 var2} G {t}
- : interp_flat_type (@inverted_expr var1) t
- -> interp_flat_type (@inverted_expr var2) t
- -> Prop
- := match t with
- | Tbase T
- => fun z1 z2 => match z1, z2 return Prop 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
- | Unit => fun _ _ => True
- | Prod A B => fun a b : interp_flat_type _ A * interp_flat_type _ B
- => and (@wff_as_expr_or_const var1 var2 G A (fst a) (fst b))
- (@wff_as_expr_or_const var1 var2 G B (snd a) (snd b))
- end.
-
-Lemma wff_interp_as_expr_or_const {var1 var2 t} {G e1 e2 v1 v2}
- (Hwf : @wff var1 var2 G t e1 e2)
- : @interp_as_expr_or_const var1 t e1 = Some v1
- -> @interp_as_expr_or_const var2 t e2 = Some v2
- -> wff_as_expr_or_const G v1 v2.
-Proof.
- induction Hwf;
- repeat first [ progress subst
- | progress inversion_option
- | progress simpl in *
- | progress cbn [wff_as_expr_or_const]
- | reflexivity
- | break_innermost_match_hyps_step
- | intro
- | match goal with
- | [ H : forall z, Some _ = Some z -> _ |- _ ] => specialize (H _ eq_refl)
- | [ |- context[match ?e with _ => _ end] ]
- => is_var e; invert_one_op e
- end
- | break_innermost_match_step
- | solve [ auto with wf ] ].
-Qed.
-
-Local Ltac pose_wff _ :=
- match goal with
- | [ H1 : _ = Some _, H2 : _ = Some _, Hwf : wff _ _ _ |- _ ]
- => pose proof (wff_interp_as_expr_or_const Hwf H1 H2); clear H1 H2
- end.
-
-Lemma Wf_SimplifyArith {convert_adc_to_sbb} {t} (e : Expr t)
- (Hwf : Wf e)
- : Wf (SimplifyArith convert_adc_to_sbb e).
-Proof.
- apply Wf_RewriteOp; [ | assumption ].
- intros ???????? Hwf'; unfold simplify_op_expr.
- repeat match goal with
- | [ H : ?T |- ?T ] => exact H
- | [ H : False |- _ ] => exfalso; assumption
- | [ |- True ] => exact I
- | [ H : false = true |- _ ] => exfalso; clear -H; discriminate
- | [ H : true = false |- _ ] => exfalso; clear -H; discriminate
- | [ H : None = Some _ |- _ ] => exfalso; clear -H; discriminate
- | [ H : Some _ = None |- _ ] => exfalso; clear -H; discriminate
- | [ H : TT = Op _ _ |- _ ] => exfalso; clear -H; discriminate
- | [ H : invert_Op ?e = None, H' : wff _ (Op ?opc _) ?e |- _ ]
- => progress (exfalso; clear -H H'; generalize dependent opc; intros; try (is_var e; destruct e))
- | [ H : invert_Op ?e = None, H' : wff _ ?e (Op ?opc _) |- _ ]
- => progress (exfalso; clear -H H'; generalize dependent opc; intros; try (is_var e; destruct e))
- | _ => progress destruct_head'_and
- | _ => progress subst
- | _ => progress destruct_head'_prod
- | _ => progress destruct_head'_sig
- | _ => progress destruct_head'_sigT
- | _ => inversion_base_type_constr_step
- | _ => inversion_wf_step_constr
- | _ => progress invert_expr_subst
- | _ => progress rewrite_eta_match_base_type_impl
- | [ H : ?x = ?x |- _ ] => clear H || (progress eliminate_hprop_eq)
- | [ H : match ?e with @const_of _ _ _ => _ = _ | _ => _ end |- _ ]
- => is_var e; destruct e
- | [ H : match ?e with @const_of _ _ _ => False | _ => _ end |- _ ]
- => is_var e; destruct e
- | [ 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 _ (Op _ _) (LetIn _ _) ] => exfalso
- | [ |- wff _ (LetIn _ _) (Op _ _) ] => exfalso
- | [ |- wff _ (Pair _ _) (LetIn _ _) ] => exfalso
- | [ |- wff _ (LetIn _ _) (Pair _ _) ] => exfalso
- | _ => pose_wff ()
- | _ => progress cbn [fst snd projT1 projT2 interp_flat_type wff_as_expr_or_const eq_rect invert_Op] in *
- | [ |- wff _ _ _ ] => constructor; intros
- | [ H : match ?e with @const_of _ _ _ => _ | _ => _ end |- _ ]
- => is_var e; destruct e
- | [ |- context[match @interp_as_expr_or_const ?var ?t ?e with _ => _ end] ]
- => destruct (@interp_as_expr_or_const var t e) eqn:?
- | [ |- context[match base_type_eq_semidec_transparent ?t1 ?t2 with _ => _ end] ]
- => destruct (base_type_eq_semidec_transparent t1 t2)
- | [ |- context[match @invert_Op ?base_type ?op ?var ?t ?e with _ => _ end] ]
- => destruct (@invert_Op base_type op var t e) eqn:?
- | [ |- context[if BinInt.Z.eqb ?x ?y then _ else _] ]
- => destruct (BinInt.Z.eqb x y) eqn:?
- | [ |- context[if BinInt.Z.ltb ?x ?y then _ else _] ]
- => destruct (BinInt.Z.ltb x y) eqn:?
- | [ |- context[match ?e with @OpConst _ _ => _ | _ => _ end] ]
- => is_var e; destruct e
- | [ |- context[match ?e with @OpConst _ _ => _ | _ => _ end] ]
- => is_var e; invert_one_op e; try exact I; break_innermost_match_step; intros
- | [ |- List.In _ _ ] => progress (simpl; auto)
- | _ => break_innermost_match_step
- end.
-Qed.
-
-Hint Resolve Wf_SimplifyArith : wf.
diff --git a/src/Compilers/Z/BinaryNotationConstants.v b/src/Compilers/Z/BinaryNotationConstants.v
deleted file mode 100644
index 5ecd088bd..000000000
--- a/src/Compilers/Z/BinaryNotationConstants.v
+++ /dev/null
@@ -1,3901 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Export Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Export bbv.WordScope.
-Require Export Crypto.Util.Notations.
-
-Notation Const x := (Op (OpConst x) TT).
-(* python:
-<<
-#!/usr/bin/env python
-from __future__ import with_statement
-import math
-
-systematic_nums \
- = (list(range(128)) +
- [2**i + j for i in range(256) for j in (1, 0, -1)])
-nums = tuple(sorted(set(systematic_nums + [
- 187,
- 189,
- 317,
- 421,
- 481,
- 569,
- 765,
- 977,
- 5311,
- 32765,
- 32766,
- 65530,
- 65531,
- 65534,
- 114687,
- 121665,
- 130307,
- 131039,
- 131062,
- 131067,
- 261167,
- 261575,
- 262131,
- 262139,
- 262142,
- 523807,
- 524101,
- 524223,
- 524262,
- 524263,
- 524267,
- 524269,
- 524271,
- 524278,
- 524279,
- 524286,
- 679935,
- 1032192,
- 1044479,
- 1047599,
- 1047811,
- 1048202,
- 1048446,
- 1048471,
- 1048511,
- 1048526,
- 1048534,
- 1048538,
- 1048549,
- 1048557,
- 1048558,
- 1048559,
- 1048573,
- 1048574,
- 1359870,
- 2060031,
- 2064384,
- 2081439,
- 2088958,
- 2094335,
- 2095198,
- 2095622,
- 2096127,
- 2096128,
- 2096583,
- 2096942,
- 2097098,
- 2097114,
- 2097118,
- 2097127,
- 2097135,
- 2097143,
- 2097150,
- 4120062,
- 4162878,
- 4188670,
- 4192254,
- 4193166,
- 4193327,
- 4193539,
- 4193735,
- 4193823,
- 4193883,
- 4193987,
- 4194115,
- 4194240,
- 4194254,
- 4194270,
- 4194271,
- 4194275,
- 4194279,
- 4194285,
- 4194286,
- 4194287,
- 4194289,
- 4194293,
- 4194299,
- 4194301,
- 4194302,
- 8323583,
- 8372224,
- 8386654,
- 8387078,
- 8387470,
- 8387646,
- 8387766,
- 8387974,
- 8388127,
- 8388187,
- 8388230,
- 8388291,
- 8388421,
- 8388491,
- 8388503,
- 8388542,
- 8388558,
- 8388570,
- 8388574,
- 8388577,
- 8388578,
- 8388581,
- 8388591,
- 8388593,
- 8388595,
- 8388598,
- 8388599,
- 8388602,
- 8388603,
- 8388604,
- 8388605,
- 8388606,
- 11599871,
- 16647166,
- 16711679,
- 16775935,
- 16776254,
- 16776374,
- 16776582,
- 16776842,
- 16776959,
- 16776982,
- 16777006,
- 16777027,
- 16777029,
- 16777111,
- 16777162,
- 16777182,
- 16777185,
- 16777186,
- 16777187,
- 16777189,
- 16777190,
- 16777191,
- 16777197,
- 16777198,
- 16777199,
- 16777201,
- 16777205,
- 16777206,
- 16777207,
- 16777210,
- 16777211,
- 16777212,
- 16777213,
- 16777214,
- 23199742,
- 33423358,
- 33551870,
- 33554054,
- 33554058,
- 33554222,
- 33554315,
- 33554370,
- 33554374,
- 33554378,
- 33554382,
- 33554394,
- 33554398,
- 33554399,
- 33554401,
- 33554402,
- 33554407,
- 33554410,
- 33554411,
- 33554413,
- 33554414,
- 33554415,
- 33554417,
- 33554422,
- 33554423,
- 33554426,
- 33554427,
- 33554428,
- 33554429,
- 33554430,
- 67107887,
- 67108630,
- 67108798,
- 67108799,
- 67108802,
- 67108814,
- 67108822,
- 67108826,
- 67108830,
- 67108833,
- 67108834,
- 67108837,
- 67108839,
- 67108843,
- 67108845,
- 67108846,
- 67108847,
- 67108849,
- 67108854,
- 67108855,
- 67108858,
- 67108859,
- 67108860,
- 67108861,
- 67108862,
- 134215774,
- 134217598,
- 134217666,
- 134217674,
- 134217678,
- 134217686,
- 134217690,
- 134217694,
- 134217697,
- 134217698,
- 134217699,
- 134217703,
- 134217709,
- 134217710,
- 134217711,
- 134217713,
- 134217718,
- 134217719,
- 134217722,
- 134217724,
- 134217725,
- 134217726,
- 268431360,
- 268435394,
- 268435398,
- 268435406,
- 268435418,
- 268435422,
- 268435426,
- 268435438,
- 268435441,
- 268435443,
- 268435445,
- 268435447,
- 268435450,
- 268435451,
- 268435452,
- 268435453,
- 268435454,
- 536869935,
- 536870882,
- 536870886,
- 536870890,
- 536870893,
- 536870894,
- 536870902,
- 536870906,
- 536870907,
- 536870908,
- 536870909,
- 536870910,
- 678152731,
- 954437177,
- 1054736383,
- 1065418751,
- 1073709055,
- 1073739870,
- 1073741786,
- 1073741814,
- 1073741818,
- 1073741821,
- 1073741822,
- 1332920885,
- 1749801491,
- 2147418110,
- 2147483631,
- 2147483642,
- 2147483644,
- 2147483646,
- 2147483647,
- 2863311531,
- 2969567231,
- 3123612579,
- 3264175145,
- 3303820997,
- 3435973837,
- 3707621341,
- 4008636143,
- 4042322161,
- 4262789119,
- 4289200127,
- 4294639615,
- 4294901759,
- 4294963199,
- 4294966319,
- 4294966531,
- 4294966727,
- 4294966815,
- 4294966875,
- 4294966979,
- 4294967107,
- 4294967109,
- 4294967179,
- 4294967191,
- 4294967262,
- 4294967263,
- 4294967265,
- 4294967267,
- 4294967269,
- 4294967271,
- 4294967275,
- 4294967277,
- 4294967279,
- 4294967281,
- 4294967283,
- 4294967285,
- 4294967287,
- 4294967291,
- 4294967293,
- 4294967294,
- 4294968273,
- 8589934559,
- 8589934567,
- 8589934575,
- 8589934587,
- 8589934590,
- 17179869134,
- 17179869174,
- 17179869182,
- 34359738341,
- 34359738349,
- 34359738355,
- 34359738366,
- 68719476682,
- 68719476707,
- 68719476710,
- 68719476727,
- 68719476733,
- 68719476734,
- 133143985199,
- 137436856320,
- 137438952991,
- 137438953285,
- 137438953355,
- 137438953454,
- 137438953470,
- 266287970398,
- 274844352511,
- 274877902847,
- 274877906919,
- 274877906927,
- 274877906933,
- 274877906939,
- 274877906941,
- 274877906942,
- 549688705022,
- 549755289600,
- 549755813783,
- 549755813838,
- 549755813854,
- 549755813855,
- 549755813866,
- 549755813869,
- 549755813878,
- 549755813882,
- 549755813886,
- 1099511627566,
- 1099511627710,
- 1099511627738,
- 1099511627759,
- 1099511627761,
- 1099511627774,
- 1425929142271,
- 2199023255518,
- 2199023255522,
- 2199023255543,
- 2199023255550,
- 4363955208191,
- 4398046510080,
- 4398046511079,
- 4398046511086,
- 4398046511095,
- 4398046511099,
- 4398046511102,
- 8727910416382,
- 8791798053935,
- 8796090925055,
- 8796093021443,
- 8796093022019,
- 8796093022158,
- 8796093022179,
- 8796093022183,
- 8796093022189,
- 8796093022190,
- 8796093022193,
- 8796093022198,
- 8796093022205,
- 8796093022206,
- 17583596107870,
- 17592181850110,
- 17592186042886,
- 17592186044038,
- 17592186044358,
- 17592186044366,
- 17592186044378,
- 17592186044399,
- 17592186044410,
- 17592186044411,
- 17592186044413,
- 17592186044414,
- 35184372088715,
- 35184372088768,
- 35184372088798,
- 35184372088822,
- 35184372088826,
- 35184372088829,
- 35184372088830,
- 70368735789055,
- 70368744177430,
- 70368744177637,
- 70368744177647,
- 70368744177651,
- 70368744177655,
- 70368744177658,
- 70368744177659,
- 70368744177660,
- 70368744177661,
- 70368744177662,
- 140737471578110,
- 140737479966719,
- 140737488354759,
- 140737488354847,
- 140737488355141,
- 140737488355274,
- 140737488355294,
- 140737488355301,
- 140737488355302,
- 140737488355303,
- 140737488355310,
- 140737488355313,
- 140737488355318,
- 140737488355319,
- 140737488355326,
- 194613558116351,
- 281453501874175,
- 281470681743359,
- 281474959933438,
- 281474976645119,
- 281474976709518,
- 281474976709694,
- 281474976710235,
- 281474976710282,
- 281474976710339,
- 281474976710469,
- 281474976710551,
- 281474976710602,
- 281474976710606,
- 281474976710625,
- 281474976710626,
- 281474976710631,
- 281474976710637,
- 281474976710638,
- 281474976710639,
- 281474976710645,
- 281474976710647,
- 281474976710653,
- 281474976710654,
- 389227116232702,
- 562907003748350,
- 562949953290238,
- 562949953420470,
- 562949953420678,
- 562949953420938,
- 562949953421102,
- 562949953421250,
- 562949953421262,
- 562949953421274,
- 562949953421278,
- 562949953421279,
- 562949953421290,
- 562949953421291,
- 562949953421293,
- 562949953421294,
- 562949953421295,
- 562949953421297,
- 562949953421303,
- 562949953421306,
- 562949953421308,
- 562949953421310,
- 1125899873288191,
- 1125899906842558,
- 1125899906842582,
- 1125899906842586,
- 1125899906842590,
- 1125899906842593,
- 1125899906842594,
- 1125899906842606,
- 1125899906842607,
- 1125899906842619,
- 1125899906842621,
- 1125899906842622,
- 1460151441686527,
- 2211942517178367,
- 2234929182146559,
- 2248776156708863,
- 2251799813160960,
- 2251799813684483,
- 2251799813685186,
- 2251799813685210,
- 2251799813685214,
- 2251799813685217,
- 2251799813685229,
- 2251799813685231,
- 2251799813685238,
- 2251799813685239,
- 2251799813685242,
- 2251799813685245,
- 2251799813685246,
- 2920302883373054,
- 4423885034356734,
- 4469858364293118,
- 4497552313417726,
- 4503595332402223,
- 4503599627368966,
- 4503599627369927,
- 4503599627370015,
- 4503599627370307,
- 4503599627370309,
- 4503599627370434,
- 4503599627370458,
- 4503599627370462,
- 4503599627370475,
- 4503599627370478,
- 4503599627370479,
- 4503599627370490,
- 4503599627370491,
- 4503599627370494,
- 9007190664804446,
- 9007199187632127,
- 9007199254739854,
- 9007199254740030,
- 9007199254740614,
- 9007199254740618,
- 9007199254740950,
- 9007199254740958,
- 9007199254740963,
- 9007199254740967,
- 9007199254740977,
- 9007199254740982,
- 9007199254740988,
- 9007199254740990,
- 18014398509481926,
- 18014398509481934,
- 18014398509481954,
- 18014398509481975,
- 18014398509481981,
- 18014398509481982,
- 36028797018963547,
- 36028797018963651,
- 36028797018963781,
- 36028797018963863,
- 36028797018963937,
- 36028797018963943,
- 36028797018963947,
- 36028797018963949,
- 36028797018963950,
- 36028797018963951,
- 36028797018963962,
- 36028797018963964,
- 36028797018963966,
- 72056494526300160,
- 72057594037927094,
- 72057594037927302,
- 72057594037927562,
- 72057594037927726,
- 72057594037927819,
- 72057594037927874,
- 72057594037927886,
- 72057594037927894,
- 72057594037927898,
- 72057594037927902,
- 72057594037927915,
- 72057594037927919,
- 72057594037927931,
- 72057594037927933,
- 72057594037927934,
- 144115188075855638,
- 144115188075855830,
- 144115188075855838,
- 144115188075855853,
- 144115188075855857,
- 144115188075855862,
- 144115188075855863,
- 144115188075855866,
- 144115188075855867,
- 144115188075855868,
- 144115188075855869,
- 144115188075855870,
- 288230376151711706,
- 288230376151711713,
- 288230376151711714,
- 288230376151711717,
- 288230376151711726,
- 288230376151711727,
- 288230376151711734,
- 288230376151711738,
- 288230376151711740,
- 288230376151711741,
- 288230376151711742,
- 576460752303423426,
- 576460752303423434,
- 576460752303423454,
- 576460752303423467,
- 576460752303423469,
- 576460752303423471,
- 576460752303423473,
- 576460752303423482,
- 576460752303423486,
- 1117984489315730401,
- 1152921504606846934,
- 1152921504606846938,
- 1152921504606846942,
- 1152921504606846946,
- 1152921504606846974,
- 2305843009213693948,
- 2305843009213693950,
- 3353953467947191203,
- 3816567739388183093,
- 4530058275181297663,
- 4575938696385134591,
- 5675921253449092805,
- 9564978408590137875,
- 9708812670373448219,
- 9963214713607832691,
- 10248191152060862009,
- 10330176681277348905,
- 10365313336655843289,
- 11907422100489763477,
- 12273715991527008853,
- 12297829382473034411,
- 12754194144713244671,
- 14757395258967641293,
- 14897608040525528621,
- 14933078535860113213,
- 14978125529935106013,
- 15580212934572586289,
- 17050145153302519317,
- 17216961135462248175,
- 17256631552825064415,
- 17361641481138401521,
- 18308539860144619519,
- 18421974275759013887,
- 18445336698825998335,
- 18446462598732840959,
- 18446726481523507199,
- 18446744065119617023,
- 18446744069414583343,
- 18446744069414584319,
- 18446744069414584320,
- 18446744069414584321,
- 18446744073709486079,
- 18446744073709550851,
- 18446744073709551047,
- 18446744073709551135,
- 18446744073709551195,
- 18446744073709551299,
- 18446744073709551427,
- 18446744073709551429,
- 18446744073709551499,
- 18446744073709551511,
- 18446744073709551583,
- 18446744073709551585,
- 18446744073709551587,
- 18446744073709551589,
- 18446744073709551591,
- 18446744073709551595,
- 18446744073709551597,
- 18446744073709551599,
- 18446744073709551601,
- 18446744073709551603,
- 18446744073709551605,
- 18446744073709551607,
- 18446744073709551611,
- 18446744073709551613,
- 18446744073709551614,
- 77371252455336267181195254,
- 77371252455336267181195262,
- 340282366841710300967557013911933812736,
-])))
-
-def log2_up(i):
- return int(math.ceil(math.log(i, 2)))
-
-def word(i, width=None):
- assert(i >= 0)
- if width is None:
- if i == 0: return word(i, width=1)
- return word(i, width=2**log2_up(log2_up(i + 1)))
- return 'WO~' + '~'.join(bin(i)[2:].rjust(width, '0'))
-
-word_formats = tuple(sorted([(n, hex(n), bin(n), word(n)) for n in nums]))
-
-def header():
- return (r"""Require Import Coq.ZArith.ZArith.
-Require Export Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Export bbv.WordScope.
-Require Export Crypto.Util.Notations.
-
-Notation Const x := (Op (OpConst x) TT).
-(""" + r"""* python:
-<<
-""" +
- open(__file__).read() +
- r""">>
- *""" + r""")""")
-
-
-def hex_nots():
- return ('\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 word_formats))
-def bin_nots():
- return ('\n'.join('''Notation "'%s'" (* %d (%s) *)\n := (Const %s).''' % (b, d, h, w)
- for d, h, b, w in word_formats))
-
-with open('HexNotationConstants.v', 'w') as f:
- f.write(header() + '\n' + hex_nots() + '\n')
-
-with open('BinaryNotationConstants.v', 'w') as f:
- f.write(header() + '\n' + bin_nots() + '\n')
->>
- *)
-Notation "'0b0'" (* 0 (0x0) *)
- := (Const WO~0).
-Notation "'0b1'" (* 1 (0x1) *)
- := (Const WO~1).
-Notation "'0b10'" (* 2 (0x2) *)
- := (Const WO~1~0).
-Notation "'0b11'" (* 3 (0x3) *)
- := (Const WO~1~1).
-Notation "'0b100'" (* 4 (0x4) *)
- := (Const WO~0~1~0~0).
-Notation "'0b101'" (* 5 (0x5) *)
- := (Const WO~0~1~0~1).
-Notation "'0b110'" (* 6 (0x6) *)
- := (Const WO~0~1~1~0).
-Notation "'0b111'" (* 7 (0x7) *)
- := (Const WO~0~1~1~1).
-Notation "'0b1000'" (* 8 (0x8) *)
- := (Const WO~1~0~0~0).
-Notation "'0b1001'" (* 9 (0x9) *)
- := (Const WO~1~0~0~1).
-Notation "'0b1010'" (* 10 (0xa) *)
- := (Const WO~1~0~1~0).
-Notation "'0b1011'" (* 11 (0xb) *)
- := (Const WO~1~0~1~1).
-Notation "'0b1100'" (* 12 (0xc) *)
- := (Const WO~1~1~0~0).
-Notation "'0b1101'" (* 13 (0xd) *)
- := (Const WO~1~1~0~1).
-Notation "'0b1110'" (* 14 (0xe) *)
- := (Const WO~1~1~1~0).
-Notation "'0b1111'" (* 15 (0xf) *)
- := (Const WO~1~1~1~1).
-Notation "'0b10000'" (* 16 (0x10) *)
- := (Const WO~0~0~0~1~0~0~0~0).
-Notation "'0b10001'" (* 17 (0x11) *)
- := (Const WO~0~0~0~1~0~0~0~1).
-Notation "'0b10010'" (* 18 (0x12) *)
- := (Const WO~0~0~0~1~0~0~1~0).
-Notation "'0b10011'" (* 19 (0x13) *)
- := (Const WO~0~0~0~1~0~0~1~1).
-Notation "'0b10100'" (* 20 (0x14) *)
- := (Const WO~0~0~0~1~0~1~0~0).
-Notation "'0b10101'" (* 21 (0x15) *)
- := (Const WO~0~0~0~1~0~1~0~1).
-Notation "'0b10110'" (* 22 (0x16) *)
- := (Const WO~0~0~0~1~0~1~1~0).
-Notation "'0b10111'" (* 23 (0x17) *)
- := (Const WO~0~0~0~1~0~1~1~1).
-Notation "'0b11000'" (* 24 (0x18) *)
- := (Const WO~0~0~0~1~1~0~0~0).
-Notation "'0b11001'" (* 25 (0x19) *)
- := (Const WO~0~0~0~1~1~0~0~1).
-Notation "'0b11010'" (* 26 (0x1a) *)
- := (Const WO~0~0~0~1~1~0~1~0).
-Notation "'0b11011'" (* 27 (0x1b) *)
- := (Const WO~0~0~0~1~1~0~1~1).
-Notation "'0b11100'" (* 28 (0x1c) *)
- := (Const WO~0~0~0~1~1~1~0~0).
-Notation "'0b11101'" (* 29 (0x1d) *)
- := (Const WO~0~0~0~1~1~1~0~1).
-Notation "'0b11110'" (* 30 (0x1e) *)
- := (Const WO~0~0~0~1~1~1~1~0).
-Notation "'0b11111'" (* 31 (0x1f) *)
- := (Const WO~0~0~0~1~1~1~1~1).
-Notation "'0b100000'" (* 32 (0x20) *)
- := (Const WO~0~0~1~0~0~0~0~0).
-Notation "'0b100001'" (* 33 (0x21) *)
- := (Const WO~0~0~1~0~0~0~0~1).
-Notation "'0b100010'" (* 34 (0x22) *)
- := (Const WO~0~0~1~0~0~0~1~0).
-Notation "'0b100011'" (* 35 (0x23) *)
- := (Const WO~0~0~1~0~0~0~1~1).
-Notation "'0b100100'" (* 36 (0x24) *)
- := (Const WO~0~0~1~0~0~1~0~0).
-Notation "'0b100101'" (* 37 (0x25) *)
- := (Const WO~0~0~1~0~0~1~0~1).
-Notation "'0b100110'" (* 38 (0x26) *)
- := (Const WO~0~0~1~0~0~1~1~0).
-Notation "'0b100111'" (* 39 (0x27) *)
- := (Const WO~0~0~1~0~0~1~1~1).
-Notation "'0b101000'" (* 40 (0x28) *)
- := (Const WO~0~0~1~0~1~0~0~0).
-Notation "'0b101001'" (* 41 (0x29) *)
- := (Const WO~0~0~1~0~1~0~0~1).
-Notation "'0b101010'" (* 42 (0x2a) *)
- := (Const WO~0~0~1~0~1~0~1~0).
-Notation "'0b101011'" (* 43 (0x2b) *)
- := (Const WO~0~0~1~0~1~0~1~1).
-Notation "'0b101100'" (* 44 (0x2c) *)
- := (Const WO~0~0~1~0~1~1~0~0).
-Notation "'0b101101'" (* 45 (0x2d) *)
- := (Const WO~0~0~1~0~1~1~0~1).
-Notation "'0b101110'" (* 46 (0x2e) *)
- := (Const WO~0~0~1~0~1~1~1~0).
-Notation "'0b101111'" (* 47 (0x2f) *)
- := (Const WO~0~0~1~0~1~1~1~1).
-Notation "'0b110000'" (* 48 (0x30) *)
- := (Const WO~0~0~1~1~0~0~0~0).
-Notation "'0b110001'" (* 49 (0x31) *)
- := (Const WO~0~0~1~1~0~0~0~1).
-Notation "'0b110010'" (* 50 (0x32) *)
- := (Const WO~0~0~1~1~0~0~1~0).
-Notation "'0b110011'" (* 51 (0x33) *)
- := (Const WO~0~0~1~1~0~0~1~1).
-Notation "'0b110100'" (* 52 (0x34) *)
- := (Const WO~0~0~1~1~0~1~0~0).
-Notation "'0b110101'" (* 53 (0x35) *)
- := (Const WO~0~0~1~1~0~1~0~1).
-Notation "'0b110110'" (* 54 (0x36) *)
- := (Const WO~0~0~1~1~0~1~1~0).
-Notation "'0b110111'" (* 55 (0x37) *)
- := (Const WO~0~0~1~1~0~1~1~1).
-Notation "'0b111000'" (* 56 (0x38) *)
- := (Const WO~0~0~1~1~1~0~0~0).
-Notation "'0b111001'" (* 57 (0x39) *)
- := (Const WO~0~0~1~1~1~0~0~1).
-Notation "'0b111010'" (* 58 (0x3a) *)
- := (Const WO~0~0~1~1~1~0~1~0).
-Notation "'0b111011'" (* 59 (0x3b) *)
- := (Const WO~0~0~1~1~1~0~1~1).
-Notation "'0b111100'" (* 60 (0x3c) *)
- := (Const WO~0~0~1~1~1~1~0~0).
-Notation "'0b111101'" (* 61 (0x3d) *)
- := (Const WO~0~0~1~1~1~1~0~1).
-Notation "'0b111110'" (* 62 (0x3e) *)
- := (Const WO~0~0~1~1~1~1~1~0).
-Notation "'0b111111'" (* 63 (0x3f) *)
- := (Const WO~0~0~1~1~1~1~1~1).
-Notation "'0b1000000'" (* 64 (0x40) *)
- := (Const WO~0~1~0~0~0~0~0~0).
-Notation "'0b1000001'" (* 65 (0x41) *)
- := (Const WO~0~1~0~0~0~0~0~1).
-Notation "'0b1000010'" (* 66 (0x42) *)
- := (Const WO~0~1~0~0~0~0~1~0).
-Notation "'0b1000011'" (* 67 (0x43) *)
- := (Const WO~0~1~0~0~0~0~1~1).
-Notation "'0b1000100'" (* 68 (0x44) *)
- := (Const WO~0~1~0~0~0~1~0~0).
-Notation "'0b1000101'" (* 69 (0x45) *)
- := (Const WO~0~1~0~0~0~1~0~1).
-Notation "'0b1000110'" (* 70 (0x46) *)
- := (Const WO~0~1~0~0~0~1~1~0).
-Notation "'0b1000111'" (* 71 (0x47) *)
- := (Const WO~0~1~0~0~0~1~1~1).
-Notation "'0b1001000'" (* 72 (0x48) *)
- := (Const WO~0~1~0~0~1~0~0~0).
-Notation "'0b1001001'" (* 73 (0x49) *)
- := (Const WO~0~1~0~0~1~0~0~1).
-Notation "'0b1001010'" (* 74 (0x4a) *)
- := (Const WO~0~1~0~0~1~0~1~0).
-Notation "'0b1001011'" (* 75 (0x4b) *)
- := (Const WO~0~1~0~0~1~0~1~1).
-Notation "'0b1001100'" (* 76 (0x4c) *)
- := (Const WO~0~1~0~0~1~1~0~0).
-Notation "'0b1001101'" (* 77 (0x4d) *)
- := (Const WO~0~1~0~0~1~1~0~1).
-Notation "'0b1001110'" (* 78 (0x4e) *)
- := (Const WO~0~1~0~0~1~1~1~0).
-Notation "'0b1001111'" (* 79 (0x4f) *)
- := (Const WO~0~1~0~0~1~1~1~1).
-Notation "'0b1010000'" (* 80 (0x50) *)
- := (Const WO~0~1~0~1~0~0~0~0).
-Notation "'0b1010001'" (* 81 (0x51) *)
- := (Const WO~0~1~0~1~0~0~0~1).
-Notation "'0b1010010'" (* 82 (0x52) *)
- := (Const WO~0~1~0~1~0~0~1~0).
-Notation "'0b1010011'" (* 83 (0x53) *)
- := (Const WO~0~1~0~1~0~0~1~1).
-Notation "'0b1010100'" (* 84 (0x54) *)
- := (Const WO~0~1~0~1~0~1~0~0).
-Notation "'0b1010101'" (* 85 (0x55) *)
- := (Const WO~0~1~0~1~0~1~0~1).
-Notation "'0b1010110'" (* 86 (0x56) *)
- := (Const WO~0~1~0~1~0~1~1~0).
-Notation "'0b1010111'" (* 87 (0x57) *)
- := (Const WO~0~1~0~1~0~1~1~1).
-Notation "'0b1011000'" (* 88 (0x58) *)
- := (Const WO~0~1~0~1~1~0~0~0).
-Notation "'0b1011001'" (* 89 (0x59) *)
- := (Const WO~0~1~0~1~1~0~0~1).
-Notation "'0b1011010'" (* 90 (0x5a) *)
- := (Const WO~0~1~0~1~1~0~1~0).
-Notation "'0b1011011'" (* 91 (0x5b) *)
- := (Const WO~0~1~0~1~1~0~1~1).
-Notation "'0b1011100'" (* 92 (0x5c) *)
- := (Const WO~0~1~0~1~1~1~0~0).
-Notation "'0b1011101'" (* 93 (0x5d) *)
- := (Const WO~0~1~0~1~1~1~0~1).
-Notation "'0b1011110'" (* 94 (0x5e) *)
- := (Const WO~0~1~0~1~1~1~1~0).
-Notation "'0b1011111'" (* 95 (0x5f) *)
- := (Const WO~0~1~0~1~1~1~1~1).
-Notation "'0b1100000'" (* 96 (0x60) *)
- := (Const WO~0~1~1~0~0~0~0~0).
-Notation "'0b1100001'" (* 97 (0x61) *)
- := (Const WO~0~1~1~0~0~0~0~1).
-Notation "'0b1100010'" (* 98 (0x62) *)
- := (Const WO~0~1~1~0~0~0~1~0).
-Notation "'0b1100011'" (* 99 (0x63) *)
- := (Const WO~0~1~1~0~0~0~1~1).
-Notation "'0b1100100'" (* 100 (0x64) *)
- := (Const WO~0~1~1~0~0~1~0~0).
-Notation "'0b1100101'" (* 101 (0x65) *)
- := (Const WO~0~1~1~0~0~1~0~1).
-Notation "'0b1100110'" (* 102 (0x66) *)
- := (Const WO~0~1~1~0~0~1~1~0).
-Notation "'0b1100111'" (* 103 (0x67) *)
- := (Const WO~0~1~1~0~0~1~1~1).
-Notation "'0b1101000'" (* 104 (0x68) *)
- := (Const WO~0~1~1~0~1~0~0~0).
-Notation "'0b1101001'" (* 105 (0x69) *)
- := (Const WO~0~1~1~0~1~0~0~1).
-Notation "'0b1101010'" (* 106 (0x6a) *)
- := (Const WO~0~1~1~0~1~0~1~0).
-Notation "'0b1101011'" (* 107 (0x6b) *)
- := (Const WO~0~1~1~0~1~0~1~1).
-Notation "'0b1101100'" (* 108 (0x6c) *)
- := (Const WO~0~1~1~0~1~1~0~0).
-Notation "'0b1101101'" (* 109 (0x6d) *)
- := (Const WO~0~1~1~0~1~1~0~1).
-Notation "'0b1101110'" (* 110 (0x6e) *)
- := (Const WO~0~1~1~0~1~1~1~0).
-Notation "'0b1101111'" (* 111 (0x6f) *)
- := (Const WO~0~1~1~0~1~1~1~1).
-Notation "'0b1110000'" (* 112 (0x70) *)
- := (Const WO~0~1~1~1~0~0~0~0).
-Notation "'0b1110001'" (* 113 (0x71) *)
- := (Const WO~0~1~1~1~0~0~0~1).
-Notation "'0b1110010'" (* 114 (0x72) *)
- := (Const WO~0~1~1~1~0~0~1~0).
-Notation "'0b1110011'" (* 115 (0x73) *)
- := (Const WO~0~1~1~1~0~0~1~1).
-Notation "'0b1110100'" (* 116 (0x74) *)
- := (Const WO~0~1~1~1~0~1~0~0).
-Notation "'0b1110101'" (* 117 (0x75) *)
- := (Const WO~0~1~1~1~0~1~0~1).
-Notation "'0b1110110'" (* 118 (0x76) *)
- := (Const WO~0~1~1~1~0~1~1~0).
-Notation "'0b1110111'" (* 119 (0x77) *)
- := (Const WO~0~1~1~1~0~1~1~1).
-Notation "'0b1111000'" (* 120 (0x78) *)
- := (Const WO~0~1~1~1~1~0~0~0).
-Notation "'0b1111001'" (* 121 (0x79) *)
- := (Const WO~0~1~1~1~1~0~0~1).
-Notation "'0b1111010'" (* 122 (0x7a) *)
- := (Const WO~0~1~1~1~1~0~1~0).
-Notation "'0b1111011'" (* 123 (0x7b) *)
- := (Const WO~0~1~1~1~1~0~1~1).
-Notation "'0b1111100'" (* 124 (0x7c) *)
- := (Const WO~0~1~1~1~1~1~0~0).
-Notation "'0b1111101'" (* 125 (0x7d) *)
- := (Const WO~0~1~1~1~1~1~0~1).
-Notation "'0b1111110'" (* 126 (0x7e) *)
- := (Const WO~0~1~1~1~1~1~1~0).
-Notation "'0b1111111'" (* 127 (0x7f) *)
- := (Const WO~0~1~1~1~1~1~1~1).
-Notation "'0b10000000'" (* 128 (0x80) *)
- := (Const WO~1~0~0~0~0~0~0~0).
-Notation "'0b10000001'" (* 129 (0x81) *)
- := (Const WO~1~0~0~0~0~0~0~1).
-Notation "'0b10111011'" (* 187 (0xbb) *)
- := (Const WO~1~0~1~1~1~0~1~1).
-Notation "'0b10111101'" (* 189 (0xbd) *)
- := (Const WO~1~0~1~1~1~1~0~1).
-Notation "'0b11111111'" (* 255 (0xff) *)
- := (Const WO~1~1~1~1~1~1~1~1).
-Notation "'0b100000000'" (* 256 (0x100) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0).
-Notation "'0b100000001'" (* 257 (0x101) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~1).
-Notation "'0b100111101'" (* 317 (0x13d) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~1~1~1~1~0~1).
-Notation "'0b110100101'" (* 421 (0x1a5) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~0~1~0~0~1~0~1).
-Notation "'0b111100001'" (* 481 (0x1e1) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~0~0~0~0~1).
-Notation "'0b111111111'" (* 511 (0x1ff) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1).
-Notation "'0b1000000000'" (* 512 (0x200) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000001'" (* 513 (0x201) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~1).
-Notation "'0b1000111001'" (* 569 (0x239) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~1~1~1~0~0~1).
-Notation "'0b1011111101'" (* 765 (0x2fd) *)
- := (Const WO~0~0~0~0~0~0~1~0~1~1~1~1~1~1~0~1).
-Notation "'0b1111010001'" (* 977 (0x3d1) *)
- := (Const WO~0~0~0~0~0~0~1~1~1~1~0~1~0~0~0~1).
-Notation "'0b1111111111'" (* 1023 (0x3ff) *)
- := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b10000000000'" (* 1024 (0x400) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000001'" (* 1025 (0x401) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111'" (* 2047 (0x7ff) *)
- := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b100000000000'" (* 2048 (0x800) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000001'" (* 2049 (0x801) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111'" (* 4095 (0xfff) *)
- := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1000000000000'" (* 4096 (0x1000) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000001'" (* 4097 (0x1001) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1010010111111'" (* 5311 (0x14bf) *)
- := (Const WO~0~0~0~1~0~1~0~0~1~0~1~1~1~1~1~1).
-Notation "'0b1111111111111'" (* 8191 (0x1fff) *)
- := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b10000000000000'" (* 8192 (0x2000) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000001'" (* 8193 (0x2001) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111'" (* 16383 (0x3fff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b100000000000000'" (* 16384 (0x4000) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000001'" (* 16385 (0x4001) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111101'" (* 32765 (0x7ffd) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1).
-Notation "'0b111111111111110'" (* 32766 (0x7ffe) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b111111111111111'" (* 32767 (0x7fff) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1000000000000000'" (* 32768 (0x8000) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000001'" (* 32769 (0x8001) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111010'" (* 65530 (0xfffa) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0).
-Notation "'0b1111111111111011'" (* 65531 (0xfffb) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1).
-Notation "'0b1111111111111110'" (* 65534 (0xfffe) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111111111111'" (* 65535 (0xffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b10000000000000000'" (* 65536 (0x10000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000001'" (* 65537 (0x10001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11011111111111111'" (* 114687 (0x1bfff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11101101101000001'" (* 121665 (0x1db41) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~0~1~1~0~1~1~0~1~0~0~0~0~0~1).
-Notation "'0b11111110100000011'" (* 130307 (0x1fd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1).
-Notation "'0b11111111111011111'" (* 131039 (0x1ffdf) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1).
-Notation "'0b11111111111110110'" (* 131062 (0x1fff6) *)
- := (Const WO~0~0~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~0~1~1~0).
-Notation "'0b11111111111111011'" (* 131067 (0x1fffb) *)
- := (Const WO~0~0~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~0~1~1).
-Notation "'0b11111111111111111'" (* 131071 (0x1ffff) *)
- := (Const WO~0~0~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).
-Notation "'0b100000000000000000'" (* 131072 (0x20000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000001'" (* 131073 (0x20001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111110000101111'" (* 261167 (0x3fc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b111111110111000111'" (* 261575 (0x3fdc7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1).
-Notation "'0b111111111111110011'" (* 262131 (0x3fff3) *)
- := (Const WO~0~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~0~0~1~1).
-Notation "'0b111111111111111011'" (* 262139 (0x3fffb) *)
- := (Const WO~0~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~0~1~1).
-Notation "'0b111111111111111110'" (* 262142 (0x3fffe) *)
- := (Const WO~0~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~0).
-Notation "'0b111111111111111111'" (* 262143 (0x3ffff) *)
- := (Const WO~0~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).
-Notation "'0b1000000000000000000'" (* 262144 (0x40000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000001'" (* 262145 (0x40001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111000011111'" (* 523807 (0x7fe1f) *)
- := (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~0~0~0~0~1~1~1~1~1).
-Notation "'0b1111111111101000101'" (* 524101 (0x7ff45) *)
- := (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~0~1~0~0~0~1~0~1).
-Notation "'0b1111111111110111111'" (* 524223 (0x7ffbf) *)
- := (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~0~1~1~1~1~1~1).
-Notation "'0b1111111111111100110'" (* 524262 (0x7ffe6) *)
- := (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~0~0~1~1~0).
-Notation "'0b1111111111111100111'" (* 524263 (0x7ffe7) *)
- := (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~0~0~1~1~1).
-Notation "'0b1111111111111101011'" (* 524267 (0x7ffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0b1111111111111101101'" (* 524269 (0x7ffed) *)
- := (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~0~1~1~0~1).
-Notation "'0b1111111111111101111'" (* 524271 (0x7ffef) *)
- := (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~0~1~1~1~1).
-Notation "'0b1111111111111110110'" (* 524278 (0x7fff6) *)
- := (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~0~1~1~0).
-Notation "'0b1111111111111110111'" (* 524279 (0x7fff7) *)
- := (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~0~1~1~1).
-Notation "'0b1111111111111111110'" (* 524286 (0x7fffe) *)
- := (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~0).
-Notation "'0b1111111111111111111'" (* 524287 (0x7ffff) *)
- := (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).
-Notation "'0b10000000000000000000'" (* 524288 (0x80000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000001'" (* 524289 (0x80001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b10100101111111111111'" (* 679935 (0xa5fff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111100000000000000'" (* 1032192 (0xfc000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b11111110111111111111'" (* 1044479 (0xfefff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111110000101111'" (* 1047599 (0xffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b11111111110100000011'" (* 1047811 (0xffd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1).
-Notation "'0b11111111111010001010'" (* 1048202 (0xffe8a) *)
- := (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~0~1~0~0~0~1~0~1~0).
-Notation "'0b11111111111101111110'" (* 1048446 (0xfff7e) *)
- := (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~0~1~1~1~1~1~1~0).
-Notation "'0b11111111111110010111'" (* 1048471 (0xfff97) *)
- := (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~0~0~1~0~1~1~1).
-Notation "'0b11111111111110111111'" (* 1048511 (0xfffbf) *)
- := (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~0~1~1~1~1~1~1).
-Notation "'0b11111111111111001110'" (* 1048526 (0xfffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0b11111111111111010110'" (* 1048534 (0xfffd6) *)
- := (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~0~1~0~1~1~0).
-Notation "'0b11111111111111011010'" (* 1048538 (0xfffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0b11111111111111100101'" (* 1048549 (0xfffe5) *)
- := (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~0~0~1~0~1).
-Notation "'0b11111111111111101101'" (* 1048557 (0xfffed) *)
- := (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~0~1~1~0~1).
-Notation "'0b11111111111111101110'" (* 1048558 (0xfffee) *)
- := (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~0~1~1~1~0).
-Notation "'0b11111111111111101111'" (* 1048559 (0xfffef) *)
- := (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~0~1~1~1~1).
-Notation "'0b11111111111111111101'" (* 1048573 (0xffffd) *)
- := (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~0~1).
-Notation "'0b11111111111111111110'" (* 1048574 (0xffffe) *)
- := (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~0).
-Notation "'0b11111111111111111111'" (* 1048575 (0xfffff) *)
- := (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).
-Notation "'0b100000000000000000000'" (* 1048576 (0x100000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000001'" (* 1048577 (0x100001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b101001011111111111110'" (* 1359870 (0x14bffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b111110110111011111111'" (* 2060031 (0x1f6eff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~0~1~1~1~1~1~1~1~1).
-Notation "'0b111111000000000000000'" (* 2064384 (0x1f8000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b111111100001010011111'" (* 2081439 (0x1fc29f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~0~0~1~1~1~1~1).
-Notation "'0b111111101111111111110'" (* 2088958 (0x1fdffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b111111111010011111111'" (* 2094335 (0x1ff4ff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~0~0~1~1~1~1~1~1~1~1).
-Notation "'0b111111111100001011110'" (* 2095198 (0x1ff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0b111111111101000000110'" (* 2095622 (0x1ffa06) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0b111111111101111111111'" (* 2096127 (0x1ffbff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b111111111110000000000'" (* 2096128 (0x1ffc00) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b111111111110111000111'" (* 2096583 (0x1ffdc7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1).
-Notation "'0b111111111111100101110'" (* 2096942 (0x1fff2e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~1~0).
-Notation "'0b111111111111111001010'" (* 2097098 (0x1fffca) *)
- := (Const WO~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~0~0~1~0~1~0).
-Notation "'0b111111111111111011010'" (* 2097114 (0x1fffda) *)
- := (Const WO~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~0~1~1~0~1~0).
-Notation "'0b111111111111111011110'" (* 2097118 (0x1fffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0b111111111111111100111'" (* 2097127 (0x1fffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0b111111111111111101111'" (* 2097135 (0x1fffef) *)
- := (Const WO~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~0~1~1~1~1).
-Notation "'0b111111111111111110111'" (* 2097143 (0x1ffff7) *)
- := (Const WO~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~0~1~1~1).
-Notation "'0b111111111111111111110'" (* 2097150 (0x1ffffe) *)
- := (Const WO~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~0).
-Notation "'0b111111111111111111111'" (* 2097151 (0x1fffff) *)
- := (Const WO~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).
-Notation "'0b1000000000000000000000'" (* 2097152 (0x200000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000001'" (* 2097153 (0x200001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111101101110111111110'" (* 4120062 (0x3eddfe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~0~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111000010100111110'" (* 4162878 (0x3f853e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~0~0~1~1~1~1~1~0).
-Notation "'0b1111111110100111111110'" (* 4188670 (0x3fe9fe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~0~0~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111111011111111110'" (* 4192254 (0x3ff7fe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111111101110001110'" (* 4193166 (0x3ffb8e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0b1111111111110000101111'" (* 4193327 (0x3ffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b1111111111110100000011'" (* 4193539 (0x3ffd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1).
-Notation "'0b1111111111110111000111'" (* 4193735 (0x3ffdc7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1).
-Notation "'0b1111111111111000011111'" (* 4193823 (0x3ffe1f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1).
-Notation "'0b1111111111111001011011'" (* 4193883 (0x3ffe5b) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1).
-Notation "'0b1111111111111011000011'" (* 4193987 (0x3ffec3) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~0~0~0~1~1).
-Notation "'0b1111111111111101000011'" (* 4194115 (0x3fff43) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~1~1).
-Notation "'0b1111111111111111000000'" (* 4194240 (0x3fffc0) *)
- := (Const WO~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~0~0~0~0~0~0).
-Notation "'0b1111111111111111001110'" (* 4194254 (0x3fffce) *)
- := (Const WO~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~0~0~1~1~1~0).
-Notation "'0b1111111111111111011110'" (* 4194270 (0x3fffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0b1111111111111111011111'" (* 4194271 (0x3fffdf) *)
- := (Const WO~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~0~1~1~1~1~1).
-Notation "'0b1111111111111111100011'" (* 4194275 (0x3fffe3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0b1111111111111111100111'" (* 4194279 (0x3fffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0b1111111111111111101101'" (* 4194285 (0x3fffed) *)
- := (Const WO~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~0~1~1~0~1).
-Notation "'0b1111111111111111101110'" (* 4194286 (0x3fffee) *)
- := (Const WO~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~0~1~1~1~0).
-Notation "'0b1111111111111111101111'" (* 4194287 (0x3fffef) *)
- := (Const WO~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~0~1~1~1~1).
-Notation "'0b1111111111111111110001'" (* 4194289 (0x3ffff1) *)
- := (Const WO~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~0~0~0~1).
-Notation "'0b1111111111111111110101'" (* 4194293 (0x3ffff5) *)
- := (Const WO~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~0~1~0~1).
-Notation "'0b1111111111111111111011'" (* 4194299 (0x3ffffb) *)
- := (Const WO~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~0~1~1).
-Notation "'0b1111111111111111111101'" (* 4194301 (0x3ffffd) *)
- := (Const WO~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~0~1).
-Notation "'0b1111111111111111111110'" (* 4194302 (0x3ffffe) *)
- := (Const WO~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~0).
-Notation "'0b1111111111111111111111'" (* 4194303 (0x3fffff) *)
- := (Const WO~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).
-Notation "'0b10000000000000000000000'" (* 4194304 (0x400000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000001'" (* 4194305 (0x400001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111110000000111111111'" (* 8323583 (0x7f01ff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111100000000000000'" (* 8372224 (0x7fc000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b11111111111100001011110'" (* 8386654 (0x7ff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0b11111111111101000000110'" (* 8387078 (0x7ffa06) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0b11111111111101110001110'" (* 8387470 (0x7ffb8e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0b11111111111110000111110'" (* 8387646 (0x7ffc3e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1~0).
-Notation "'0b11111111111110010110110'" (* 8387766 (0x7ffcb6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1~0).
-Notation "'0b11111111111110110000110'" (* 8387974 (0x7ffd86) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~0~0~0~1~1~0).
-Notation "'0b11111111111111000011111'" (* 8388127 (0x7ffe1f) *)
- := (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~0~0~0~0~1~1~1~1~1).
-Notation "'0b11111111111111001011011'" (* 8388187 (0x7ffe5b) *)
- := (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~0~0~1~0~1~1~0~1~1).
-Notation "'0b11111111111111010000110'" (* 8388230 (0x7ffe86) *)
- := (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~0~1~0~0~0~0~1~1~0).
-Notation "'0b11111111111111011000011'" (* 8388291 (0x7ffec3) *)
- := (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~0~1~1~0~0~0~0~1~1).
-Notation "'0b11111111111111101000101'" (* 8388421 (0x7fff45) *)
- := (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~0~1~0~0~0~1~0~1).
-Notation "'0b11111111111111110001011'" (* 8388491 (0x7fff8b) *)
- := (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~0~0~0~1~0~1~1).
-Notation "'0b11111111111111110010111'" (* 8388503 (0x7fff97) *)
- := (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~0~0~1~0~1~1~1).
-Notation "'0b11111111111111110111110'" (* 8388542 (0x7fffbe) *)
- := (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~0~1~1~1~1~1~0).
-Notation "'0b11111111111111111001110'" (* 8388558 (0x7fffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0b11111111111111111011010'" (* 8388570 (0x7fffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0b11111111111111111011110'" (* 8388574 (0x7fffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0b11111111111111111100001'" (* 8388577 (0x7fffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0b11111111111111111100010'" (* 8388578 (0x7fffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0b11111111111111111100101'" (* 8388581 (0x7fffe5) *)
- := (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~0~0~1~0~1).
-Notation "'0b11111111111111111101111'" (* 8388591 (0x7fffef) *)
- := (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~0~1~1~1~1).
-Notation "'0b11111111111111111110001'" (* 8388593 (0x7ffff1) *)
- := (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~0~0~0~1).
-Notation "'0b11111111111111111110011'" (* 8388595 (0x7ffff3) *)
- := (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~0~0~1~1).
-Notation "'0b11111111111111111110110'" (* 8388598 (0x7ffff6) *)
- := (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~0~1~1~0).
-Notation "'0b11111111111111111110111'" (* 8388599 (0x7ffff7) *)
- := (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~0~1~1~1).
-Notation "'0b11111111111111111111010'" (* 8388602 (0x7ffffa) *)
- := (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~0~1~0).
-Notation "'0b11111111111111111111011'" (* 8388603 (0x7ffffb) *)
- := (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~0~1~1).
-Notation "'0b11111111111111111111100'" (* 8388604 (0x7ffffc) *)
- := (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~0~0).
-Notation "'0b11111111111111111111101'" (* 8388605 (0x7ffffd) *)
- := (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~0~1).
-Notation "'0b11111111111111111111110'" (* 8388606 (0x7ffffe) *)
- := (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~0).
-Notation "'0b11111111111111111111111'" (* 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 "'0b100000000000000000000000'" (* 8388608 (0x800000) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000001'" (* 8388609 (0x800001) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b101100001111111111111111'" (* 11599871 (0xb0ffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~1~1~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b111111100000001111111110'" (* 16647166 (0xfe03fe) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b111111101111111111111111'" (* 16711679 (0xfeffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b111111111111101011111111'" (* 16775935 (0xfffaff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~1~1~1~1~1~1~1~1).
-Notation "'0b111111111111110000111110'" (* 16776254 (0xfffc3e) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1~0).
-Notation "'0b111111111111110010110110'" (* 16776374 (0xfffcb6) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1~0).
-Notation "'0b111111111111110110000110'" (* 16776582 (0xfffd86) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~0~0~0~1~1~0).
-Notation "'0b111111111111111010001010'" (* 16776842 (0xfffe8a) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~1~0~1~0).
-Notation "'0b111111111111111011111111'" (* 16776959 (0xfffeff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1).
-Notation "'0b111111111111111100010110'" (* 16776982 (0xffff16) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~1~0~1~1~0).
-Notation "'0b111111111111111100101110'" (* 16777006 (0xffff2e) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~1~0).
-Notation "'0b111111111111111101000011'" (* 16777027 (0xffff43) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~1~1).
-Notation "'0b111111111111111101000101'" (* 16777029 (0xffff45) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~1~0~1).
-Notation "'0b111111111111111110010111'" (* 16777111 (0xffff97) *)
- := (Const WO~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~0~0~1~0~1~1~1).
-Notation "'0b111111111111111111001010'" (* 16777162 (0xffffca) *)
- := (Const WO~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~0~0~1~0~1~0).
-Notation "'0b111111111111111111011110'" (* 16777182 (0xffffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0b111111111111111111100001'" (* 16777185 (0xffffe1) *)
- := (Const WO~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~0~0~0~0~1).
-Notation "'0b111111111111111111100010'" (* 16777186 (0xffffe2) *)
- := (Const WO~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~0~0~0~1~0).
-Notation "'0b111111111111111111100011'" (* 16777187 (0xffffe3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0b111111111111111111100101'" (* 16777189 (0xffffe5) *)
- := (Const WO~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~0~0~1~0~1).
-Notation "'0b111111111111111111100110'" (* 16777190 (0xffffe6) *)
- := (Const WO~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~0~0~1~1~0).
-Notation "'0b111111111111111111100111'" (* 16777191 (0xffffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0b111111111111111111101101'" (* 16777197 (0xffffed) *)
- := (Const WO~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~0~1~1~0~1).
-Notation "'0b111111111111111111101110'" (* 16777198 (0xffffee) *)
- := (Const WO~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~0~1~1~1~0).
-Notation "'0b111111111111111111101111'" (* 16777199 (0xffffef) *)
- := (Const WO~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~0~1~1~1~1).
-Notation "'0b111111111111111111110001'" (* 16777201 (0xfffff1) *)
- := (Const WO~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~0~0~0~1).
-Notation "'0b111111111111111111110101'" (* 16777205 (0xfffff5) *)
- := (Const WO~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~0~1~0~1).
-Notation "'0b111111111111111111110110'" (* 16777206 (0xfffff6) *)
- := (Const WO~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~0~1~1~0).
-Notation "'0b111111111111111111110111'" (* 16777207 (0xfffff7) *)
- := (Const WO~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~0~1~1~1).
-Notation "'0b111111111111111111111010'" (* 16777210 (0xfffffa) *)
- := (Const WO~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~0~1~0).
-Notation "'0b111111111111111111111011'" (* 16777211 (0xfffffb) *)
- := (Const WO~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~0~1~1).
-Notation "'0b111111111111111111111100'" (* 16777212 (0xfffffc) *)
- := (Const WO~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~0~0).
-Notation "'0b111111111111111111111101'" (* 16777213 (0xfffffd) *)
- := (Const WO~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~0~1).
-Notation "'0b111111111111111111111110'" (* 16777214 (0xfffffe) *)
- := (Const WO~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~0).
-Notation "'0b111111111111111111111111'" (* 16777215 (0xffffff) *)
- := (Const WO~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).
-Notation "'0b1000000000000000000000000'" (* 16777216 (0x1000000) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000001'" (* 16777217 (0x1000001) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1011000011111111111111110'" (* 23199742 (0x161fffe) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~1~1~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111011111111111111110'" (* 33423358 (0x1fdfffe) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111111111010111111110'" (* 33551870 (0x1fff5fe) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111111111111010000110'" (* 33554054 (0x1fffe86) *)
- := (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~0~1~0~0~0~0~1~1~0).
-Notation "'0b1111111111111111010001010'" (* 33554058 (0x1fffe8a) *)
- := (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~0~1~0~0~0~1~0~1~0).
-Notation "'0b1111111111111111100101110'" (* 33554222 (0x1ffff2e) *)
- := (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~0~0~1~0~1~1~1~0).
-Notation "'0b1111111111111111110001011'" (* 33554315 (0x1ffff8b) *)
- := (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~0~0~0~1~0~1~1).
-Notation "'0b1111111111111111111000010'" (* 33554370 (0x1ffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0b1111111111111111111000110'" (* 33554374 (0x1ffffc6) *)
- := (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~0~0~0~1~1~0).
-Notation "'0b1111111111111111111001010'" (* 33554378 (0x1ffffca) *)
- := (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~0~0~1~0~1~0).
-Notation "'0b1111111111111111111001110'" (* 33554382 (0x1ffffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0b1111111111111111111011010'" (* 33554394 (0x1ffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0b1111111111111111111011110'" (* 33554398 (0x1ffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0b1111111111111111111011111'" (* 33554399 (0x1ffffdf) *)
- := (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~0~1~1~1~1~1).
-Notation "'0b1111111111111111111100001'" (* 33554401 (0x1ffffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0b1111111111111111111100010'" (* 33554402 (0x1ffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0b1111111111111111111100111'" (* 33554407 (0x1ffffe7) *)
- := (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~0~0~1~1~1).
-Notation "'0b1111111111111111111101010'" (* 33554410 (0x1ffffea) *)
- := (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~0~1~0~1~0).
-Notation "'0b1111111111111111111101011'" (* 33554411 (0x1ffffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0b1111111111111111111101101'" (* 33554413 (0x1ffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0b1111111111111111111101110'" (* 33554414 (0x1ffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0b1111111111111111111101111'" (* 33554415 (0x1ffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0b1111111111111111111110001'" (* 33554417 (0x1fffff1) *)
- := (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~0~0~0~1).
-Notation "'0b1111111111111111111110110'" (* 33554422 (0x1fffff6) *)
- := (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~0~1~1~0).
-Notation "'0b1111111111111111111110111'" (* 33554423 (0x1fffff7) *)
- := (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~0~1~1~1).
-Notation "'0b1111111111111111111111010'" (* 33554426 (0x1fffffa) *)
- := (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~0~1~0).
-Notation "'0b1111111111111111111111011'" (* 33554427 (0x1fffffb) *)
- := (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~0~1~1).
-Notation "'0b1111111111111111111111100'" (* 33554428 (0x1fffffc) *)
- := (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~0~0).
-Notation "'0b1111111111111111111111101'" (* 33554429 (0x1fffffd) *)
- := (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~0~1).
-Notation "'0b1111111111111111111111110'" (* 33554430 (0x1fffffe) *)
- := (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~0).
-Notation "'0b1111111111111111111111111'" (* 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 "'0b10000000000000000000000000'" (* 33554432 (0x2000000) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000001'" (* 33554433 (0x2000001) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111110000101111'" (* 67107887 (0x3fffc2f) *)
- := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b11111111111111111100010110'" (* 67108630 (0x3ffff16) *)
- := (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~0~0~0~1~0~1~1~0).
-Notation "'0b11111111111111111110111110'" (* 67108798 (0x3ffffbe) *)
- := (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~0~1~1~1~1~1~0).
-Notation "'0b11111111111111111110111111'" (* 67108799 (0x3ffffbf) *)
- := (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~0~1~1~1~1~1~1).
-Notation "'0b11111111111111111111000010'" (* 67108802 (0x3ffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0b11111111111111111111001110'" (* 67108814 (0x3ffffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0b11111111111111111111010110'" (* 67108822 (0x3ffffd6) *)
- := (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~0~1~0~1~1~0).
-Notation "'0b11111111111111111111011010'" (* 67108826 (0x3ffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0b11111111111111111111011110'" (* 67108830 (0x3ffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0b11111111111111111111100001'" (* 67108833 (0x3ffffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0b11111111111111111111100010'" (* 67108834 (0x3ffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0b11111111111111111111100101'" (* 67108837 (0x3ffffe5) *)
- := (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~0~0~1~0~1).
-Notation "'0b11111111111111111111100111'" (* 67108839 (0x3ffffe7) *)
- := (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~0~0~1~1~1).
-Notation "'0b11111111111111111111101011'" (* 67108843 (0x3ffffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0b11111111111111111111101101'" (* 67108845 (0x3ffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0b11111111111111111111101110'" (* 67108846 (0x3ffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0b11111111111111111111101111'" (* 67108847 (0x3ffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0b11111111111111111111110001'" (* 67108849 (0x3fffff1) *)
- := (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~0~0~0~1).
-Notation "'0b11111111111111111111110110'" (* 67108854 (0x3fffff6) *)
- := (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~0~1~1~0).
-Notation "'0b11111111111111111111110111'" (* 67108855 (0x3fffff7) *)
- := (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~0~1~1~1).
-Notation "'0b11111111111111111111111010'" (* 67108858 (0x3fffffa) *)
- := (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~0~1~0).
-Notation "'0b11111111111111111111111011'" (* 67108859 (0x3fffffb) *)
- := (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~0~1~1).
-Notation "'0b11111111111111111111111100'" (* 67108860 (0x3fffffc) *)
- := (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~0~0).
-Notation "'0b11111111111111111111111101'" (* 67108861 (0x3fffffd) *)
- := (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~0~1).
-Notation "'0b11111111111111111111111110'" (* 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 "'0b11111111111111111111111111'" (* 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 "'0b100000000000000000000000000'" (* 67108864 (0x4000000) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000001'" (* 67108865 (0x4000001) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111100001011110'" (* 134215774 (0x7fff85e) *)
- := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0b111111111111111111101111110'" (* 134217598 (0x7ffff7e) *)
- := (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~0~1~1~1~1~1~1~0).
-Notation "'0b111111111111111111111000010'" (* 134217666 (0x7ffffc2) *)
- := (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~0~0~0~1~0).
-Notation "'0b111111111111111111111001010'" (* 134217674 (0x7ffffca) *)
- := (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~0~1~0~1~0).
-Notation "'0b111111111111111111111001110'" (* 134217678 (0x7ffffce) *)
- := (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~0~1~1~1~0).
-Notation "'0b111111111111111111111010110'" (* 134217686 (0x7ffffd6) *)
- := (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~0~1~1~0).
-Notation "'0b111111111111111111111011010'" (* 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 "'0b111111111111111111111011110'" (* 134217694 (0x7ffffde) *)
- := (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~1~1~0).
-Notation "'0b111111111111111111111100001'" (* 134217697 (0x7ffffe1) *)
- := (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~0~0~0~1).
-Notation "'0b111111111111111111111100010'" (* 134217698 (0x7ffffe2) *)
- := (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~0~0~1~0).
-Notation "'0b111111111111111111111100011'" (* 134217699 (0x7ffffe3) *)
- := (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~0~0~1~1).
-Notation "'0b111111111111111111111100111'" (* 134217703 (0x7ffffe7) *)
- := (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~0~1~1~1).
-Notation "'0b111111111111111111111101101'" (* 134217709 (0x7ffffed) *)
- := (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~0~1).
-Notation "'0b111111111111111111111101110'" (* 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 "'0b111111111111111111111101111'" (* 134217711 (0x7ffffef) *)
- := (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~1).
-Notation "'0b111111111111111111111110001'" (* 134217713 (0x7fffff1) *)
- := (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~0~0~0~1).
-Notation "'0b111111111111111111111110110'" (* 134217718 (0x7fffff6) *)
- := (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~0~1~1~0).
-Notation "'0b111111111111111111111110111'" (* 134217719 (0x7fffff7) *)
- := (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~0~1~1~1).
-Notation "'0b111111111111111111111111010'" (* 134217722 (0x7fffffa) *)
- := (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~0~1~0).
-Notation "'0b111111111111111111111111100'" (* 134217724 (0x7fffffc) *)
- := (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~0~0).
-Notation "'0b111111111111111111111111101'" (* 134217725 (0x7fffffd) *)
- := (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~0~1).
-Notation "'0b111111111111111111111111110'" (* 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 "'0b111111111111111111111111111'" (* 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 "'0b1000000000000000000000000000'" (* 134217728 (0x8000000) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000001'" (* 134217729 (0x8000001) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111000000000000'" (* 268431360 (0xffff000) *)
- := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1111111111111111111111000010'" (* 268435394 (0xfffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0b1111111111111111111111000110'" (* 268435398 (0xfffffc6) *)
- := (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~0~0~0~1~1~0).
-Notation "'0b1111111111111111111111001110'" (* 268435406 (0xfffffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0b1111111111111111111111011010'" (* 268435418 (0xfffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0b1111111111111111111111011110'" (* 268435422 (0xfffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0b1111111111111111111111100010'" (* 268435426 (0xfffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0b1111111111111111111111101110'" (* 268435438 (0xfffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0b1111111111111111111111110001'" (* 268435441 (0xffffff1) *)
- := (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~0~0~0~1).
-Notation "'0b1111111111111111111111110011'" (* 268435443 (0xffffff3) *)
- := (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~0~0~1~1).
-Notation "'0b1111111111111111111111110101'" (* 268435445 (0xffffff5) *)
- := (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~0~1~0~1).
-Notation "'0b1111111111111111111111110111'" (* 268435447 (0xffffff7) *)
- := (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~0~1~1~1).
-Notation "'0b1111111111111111111111111010'" (* 268435450 (0xffffffa) *)
- := (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~0~1~0).
-Notation "'0b1111111111111111111111111011'" (* 268435451 (0xffffffb) *)
- := (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~0~1~1).
-Notation "'0b1111111111111111111111111100'" (* 268435452 (0xffffffc) *)
- := (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~0~0).
-Notation "'0b1111111111111111111111111101'" (* 268435453 (0xffffffd) *)
- := (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~0~1).
-Notation "'0b1111111111111111111111111110'" (* 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 "'0b1111111111111111111111111111'" (* 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 "'0b10000000000000000000000000000'" (* 268435456 (0x10000000) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000001'" (* 268435457 (0x10000001) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111110000101111'" (* 536869935 (0x1ffffc2f) *)
- := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b11111111111111111111111100010'" (* 536870882 (0x1fffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0b11111111111111111111111100110'" (* 536870886 (0x1fffffe6) *)
- := (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~0~0~1~1~0).
-Notation "'0b11111111111111111111111101010'" (* 536870890 (0x1fffffea) *)
- := (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~0~1~0~1~0).
-Notation "'0b11111111111111111111111101101'" (* 536870893 (0x1fffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0b11111111111111111111111101110'" (* 536870894 (0x1fffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0b11111111111111111111111110110'" (* 536870902 (0x1ffffff6) *)
- := (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~0~1~1~0).
-Notation "'0b11111111111111111111111111010'" (* 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 "'0b11111111111111111111111111011'" (* 536870907 (0x1ffffffb) *)
- := (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~1).
-Notation "'0b11111111111111111111111111100'" (* 536870908 (0x1ffffffc) *)
- := (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~0~0).
-Notation "'0b11111111111111111111111111101'" (* 536870909 (0x1ffffffd) *)
- := (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~0~1).
-Notation "'0b11111111111111111111111111110'" (* 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 "'0b11111111111111111111111111111'" (* 536870911 (0x1fffffff) *)
- := (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~1).
-Notation "'0b100000000000000000000000000000'" (* 536870912 (0x20000000) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000001'" (* 536870913 (0x20000001) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b101000011010111100101000011011'" (* 678152731 (0x286bca1b) *)
- := (Const WO~0~0~1~0~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~1).
-Notation "'0b111000111000111000111000111001'" (* 954437177 (0x38e38e39) *)
- := (Const WO~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~1).
-Notation "'0b111110110111011111111111111111'" (* 1054736383 (0x3eddffff) *)
- := (Const WO~0~0~1~1~1~1~1~0~1~1~0~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b111111100000001111111111111111'" (* 1065418751 (0x3f80ffff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b111111111111110111111111111111'" (* 1073709055 (0x3fff7fff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b111111111111111111100001011110'" (* 1073739870 (0x3ffff85e) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0b111111111111111111111111011010'" (* 1073741786 (0x3fffffda) *)
- := (Const WO~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~0~1~1~0~1~0).
-Notation "'0b111111111111111111111111110110'" (* 1073741814 (0x3ffffff6) *)
- := (Const WO~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~1~0).
-Notation "'0b111111111111111111111111111010'" (* 1073741818 (0x3ffffffa) *)
- := (Const WO~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~1~0).
-Notation "'0b111111111111111111111111111101'" (* 1073741821 (0x3ffffffd) *)
- := (Const WO~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~1).
-Notation "'0b111111111111111111111111111110'" (* 1073741822 (0x3ffffffe) *)
- := (Const WO~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~0).
-Notation "'0b111111111111111111111111111111'" (* 1073741823 (0x3fffffff) *)
- := (Const WO~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).
-Notation "'0b1000000000000000000000000000000'" (* 1073741824 (0x40000000) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000001'" (* 1073741825 (0x40000001) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1001111011100101100001000110101'" (* 1332920885 (0x4f72c235) *)
- := (Const WO~0~1~0~0~1~1~1~1~0~1~1~1~0~0~1~0~1~1~0~0~0~0~1~0~0~0~1~1~0~1~0~1).
-Notation "'0b1101000010010111101101000010011'" (* 1749801491 (0x684bda13) *)
- := (Const WO~0~1~1~0~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~1).
-Notation "'0b1111111111111101111111111111110'" (* 2147418110 (0x7ffefffe) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111111111111111111111101111'" (* 2147483631 (0x7fffffef) *)
- := (Const WO~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~1~1~1).
-Notation "'0b1111111111111111111111111111010'" (* 2147483642 (0x7ffffffa) *)
- := (Const WO~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~1~0).
-Notation "'0b1111111111111111111111111111100'" (* 2147483644 (0x7ffffffc) *)
- := (Const WO~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~0~0).
-Notation "'0b1111111111111111111111111111110'" (* 2147483646 (0x7ffffffe) *)
- := (Const WO~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~0).
-Notation "'0b1111111111111111111111111111111'" (* 2147483647 (0x7fffffff) *)
- := (Const WO~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).
-Notation "'0b10000000000000000000000000000000'" (* 2147483648 (0x80000000) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000001'" (* 2147483649 (0x80000001) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b10101010101010101010101010101011'" (* 2863311531 (0xaaaaaaab) *)
- := (Const WO~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~1).
-Notation "'0b10110000111111111111111111111111'" (* 2969567231 (0xb0ffffff) *)
- := (Const WO~1~0~1~1~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).
-Notation "'0b10111010001011101000101110100011'" (* 3123612579 (0xba2e8ba3) *)
- := (Const WO~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~1).
-Notation "'0b11000010100011110101110000101001'" (* 3264175145 (0xc28f5c29) *)
- := (Const WO~1~1~0~0~0~0~1~0~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~1).
-Notation "'0b11000100111011000100111011000101'" (* 3303820997 (0xc4ec4ec5) *)
- := (Const WO~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~1).
-Notation "'0b11001100110011001100110011001101'" (* 3435973837 (0xcccccccd) *)
- := (Const WO~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~1).
-Notation "'0b11011100111111011100111111011101'" (* 3707621341 (0xdcfdcfdd) *)
- := (Const WO~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~1).
-Notation "'0b11101110111011101110111011101111'" (* 4008636143 (0xeeeeeeef) *)
- := (Const WO~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~1).
-Notation "'0b11110000111100001111000011110001'" (* 4042322161 (0xf0f0f0f1) *)
- := (Const WO~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~1).
-Notation "'0b11111110000101001111111111111111'" (* 4262789119 (0xfe14ffff) *)
- := (Const WO~1~1~1~1~1~1~1~0~0~0~0~1~0~1~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111101001111111111111111111'" (* 4289200127 (0xffa7ffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~0~1~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111111110101111111111111111'" (* 4294639615 (0xfffaffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111111111101111111111111111'" (* 4294901759 (0xfffeffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111111111111110111111111111'" (* 4294963199 (0xffffefff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111111111111111110000101111'" (* 4294966319 (0xfffffc2f) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b11111111111111111111110100000011'" (* 4294966531 (0xfffffd03) *)
- := (Const WO~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~0~0~0~0~0~1~1).
-Notation "'0b11111111111111111111110111000111'" (* 4294966727 (0xfffffdc7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0b11111111111111111111111000011111'" (* 4294966815 (0xfffffe1f) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1).
-Notation "'0b11111111111111111111111001011011'" (* 4294966875 (0xfffffe5b) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1).
-Notation "'0b11111111111111111111111011000011'" (* 4294966979 (0xfffffec3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0b11111111111111111111111101000011'" (* 4294967107 (0xffffff43) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0b11111111111111111111111101000101'" (* 4294967109 (0xffffff45) *)
- := (Const WO~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~0~0~1~0~1).
-Notation "'0b11111111111111111111111110001011'" (* 4294967179 (0xffffff8b) *)
- := (Const WO~1~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~0~0~1~0~1~1).
-Notation "'0b11111111111111111111111110010111'" (* 4294967191 (0xffffff97) *)
- := (Const WO~1~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~0~1~0~1~1~1).
-Notation "'0b11111111111111111111111111011110'" (* 4294967262 (0xffffffde) *)
- := (Const WO~1~1~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~1~1~0).
-Notation "'0b11111111111111111111111111011111'" (* 4294967263 (0xffffffdf) *)
- := (Const WO~1~1~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~1~1~1).
-Notation "'0b11111111111111111111111111100001'" (* 4294967265 (0xffffffe1) *)
- := (Const WO~1~1~1~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~0~0~0~1).
-Notation "'0b11111111111111111111111111100011'" (* 4294967267 (0xffffffe3) *)
- := (Const WO~1~1~1~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~0~0~1~1).
-Notation "'0b11111111111111111111111111100101'" (* 4294967269 (0xffffffe5) *)
- := (Const WO~1~1~1~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~0~1~0~1).
-Notation "'0b11111111111111111111111111100111'" (* 4294967271 (0xffffffe7) *)
- := (Const WO~1~1~1~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~0~1~1~1).
-Notation "'0b11111111111111111111111111101011'" (* 4294967275 (0xffffffeb) *)
- := (Const WO~1~1~1~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~1~1).
-Notation "'0b11111111111111111111111111101101'" (* 4294967277 (0xffffffed) *)
- := (Const WO~1~1~1~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).
-Notation "'0b11111111111111111111111111101111'" (* 4294967279 (0xffffffef) *)
- := (Const WO~1~1~1~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~1~1).
-Notation "'0b11111111111111111111111111110001'" (* 4294967281 (0xfffffff1) *)
- := (Const WO~1~1~1~1~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~0~0~1).
-Notation "'0b11111111111111111111111111110011'" (* 4294967283 (0xfffffff3) *)
- := (Const WO~1~1~1~1~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~0~1~1).
-Notation "'0b11111111111111111111111111110101'" (* 4294967285 (0xfffffff5) *)
- := (Const WO~1~1~1~1~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~1).
-Notation "'0b11111111111111111111111111110111'" (* 4294967287 (0xfffffff7) *)
- := (Const WO~1~1~1~1~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~1).
-Notation "'0b11111111111111111111111111111011'" (* 4294967291 (0xfffffffb) *)
- := (Const WO~1~1~1~1~1~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).
-Notation "'0b11111111111111111111111111111101'" (* 4294967293 (0xfffffffd) *)
- := (Const WO~1~1~1~1~1~1~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).
-Notation "'0b11111111111111111111111111111110'" (* 4294967294 (0xfffffffe) *)
- := (Const WO~1~1~1~1~1~1~1~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 "'0b11111111111111111111111111111111'" (* 4294967295 (0xffffffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000'" (* 4294967296 (0x100000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000001'" (* 4294967297 (0x100000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b100000000000000000000001111010001'" (* 4294968273 (0x1000003d1) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~0~1~0~0~0~1).
-Notation "'0b111111111111111111111111111011111'" (* 8589934559 (0x1ffffffdf) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1).
-Notation "'0b111111111111111111111111111100111'" (* 8589934567 (0x1ffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0b111111111111111111111111111101111'" (* 8589934575 (0x1ffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0b111111111111111111111111111111011'" (* 8589934587 (0x1fffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0b111111111111111111111111111111110'" (* 8589934590 (0x1fffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b111111111111111111111111111111111'" (* 8589934591 (0x1ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b1000000000000000000000000000000000'" (* 8589934592 (0x200000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000001'" (* 8589934593 (0x200000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111001110'" (* 17179869134 (0x3ffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111110110'" (* 17179869174 (0x3fffffff6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0b1111111111111111111111111111111110'" (* 17179869182 (0x3fffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b1111111111111111111111111111111111'" (* 17179869183 (0x3ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b10000000000000000000000000000000000'" (* 17179869184 (0x400000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000001'" (* 17179869185 (0x400000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111100101'" (* 34359738341 (0x7ffffffe5) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1).
-Notation "'0b11111111111111111111111111111101101'" (* 34359738349 (0x7ffffffed) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0~1).
-Notation "'0b11111111111111111111111111111110011'" (* 34359738355 (0x7fffffff3) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1).
-Notation "'0b11111111111111111111111111111111110'" (* 34359738366 (0x7fffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b11111111111111111111111111111111111'" (* 34359738367 (0x7ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b100000000000000000000000000000000000'" (* 34359738368 (0x800000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000001'" (* 34359738369 (0x800000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111001010'" (* 68719476682 (0xfffffffca) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1~0).
-Notation "'0b111111111111111111111111111111100011'" (* 68719476707 (0xfffffffe3) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1~1).
-Notation "'0b111111111111111111111111111111100110'" (* 68719476710 (0xfffffffe6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~0).
-Notation "'0b111111111111111111111111111111110111'" (* 68719476727 (0xffffffff7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1).
-Notation "'0b111111111111111111111111111111111101'" (* 68719476733 (0xffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1).
-Notation "'0b111111111111111111111111111111111110'" (* 68719476734 (0xffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b111111111111111111111111111111111111'" (* 68719476735 (0xfffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b1000000000000000000000000000000000000'" (* 68719476736 (0x1000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000001'" (* 68719476737 (0x1000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111011111111111111111111110000101111'" (* 133143985199 (0x1efffffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b1111111111111111000000000000000000000'" (* 137436856320 (0x1fffe00000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1111111111111111111111111111000011111'" (* 137438952991 (0x1ffffffe1f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~0~1~1~1~1~1).
-Notation "'0b1111111111111111111111111111101000101'" (* 137438953285 (0x1fffffff45) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0~0~0~1~0~1).
-Notation "'0b1111111111111111111111111111110001011'" (* 137438953355 (0x1fffffff8b) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1~0~1~1).
-Notation "'0b1111111111111111111111111111111101110'" (* 137438953454 (0x1fffffffee) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111110'" (* 137438953470 (0x1ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b1111111111111111111111111111111111111'" (* 137438953471 (0x1fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b10000000000000000000000000000000000000'" (* 137438953472 (0x2000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000001'" (* 137438953473 (0x2000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11110111111111111111111111100001011110'" (* 266287970398 (0x3dfffff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0b11111111111101111111111111111111111111'" (* 274844352511 (0x3ffdffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~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 "'0b11111111111111111111111110111111111111'" (* 274877902847 (0x3fffffefff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b11111111111111111111111111111111100111'" (* 274877906919 (0x3fffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0b11111111111111111111111111111111101111'" (* 274877906927 (0x3fffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0b11111111111111111111111111111111110101'" (* 274877906933 (0x3ffffffff5) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0~1).
-Notation "'0b11111111111111111111111111111111111011'" (* 274877906939 (0x3ffffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0b11111111111111111111111111111111111101'" (* 274877906941 (0x3ffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1).
-Notation "'0b11111111111111111111111111111111111110'" (* 274877906942 (0x3ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b11111111111111111111111111111111111111'" (* 274877906943 (0x3fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b100000000000000000000000000000000000000'" (* 274877906944 (0x4000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000001'" (* 274877906945 (0x4000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111011111111111111111111111110'" (* 549688705022 (0x7ffbfffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~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 "'0b111111111111111111110000000000000000000'" (* 549755289600 (0x7ffff80000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b111111111111111111111111111111110010111'" (* 549755813783 (0x7fffffff97) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1~1~1).
-Notation "'0b111111111111111111111111111111111001110'" (* 549755813838 (0x7fffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0b111111111111111111111111111111111011110'" (* 549755813854 (0x7fffffffde) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0b111111111111111111111111111111111011111'" (* 549755813855 (0x7fffffffdf) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1).
-Notation "'0b111111111111111111111111111111111101010'" (* 549755813866 (0x7fffffffea) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0~1~0).
-Notation "'0b111111111111111111111111111111111101101'" (* 549755813869 (0x7fffffffed) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0~1).
-Notation "'0b111111111111111111111111111111111110110'" (* 549755813878 (0x7ffffffff6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0b111111111111111111111111111111111111010'" (* 549755813882 (0x7ffffffffa) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0).
-Notation "'0b111111111111111111111111111111111111110'" (* 549755813886 (0x7ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b111111111111111111111111111111111111111'" (* 549755813887 (0x7fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b1000000000000000000000000000000000000000'" (* 549755813888 (0x8000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000001'" (* 549755813889 (0x8000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111100101110'" (* 1099511627566 (0xffffffff2e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111110111110'" (* 1099511627710 (0xffffffffbe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111011010'" (* 1099511627738 (0xffffffffda) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0~1~0).
-Notation "'0b1111111111111111111111111111111111101111'" (* 1099511627759 (0xffffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111110001'" (* 1099511627761 (0xfffffffff1) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111110'" (* 1099511627774 (0xfffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b1111111111111111111111111111111111111111'" (* 1099511627775 (0xffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b10000000000000000000000000000000000000000'" (* 1099511627776 (0x10000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000001'" (* 1099511627777 (0x10000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b10100101111111111111111111111111111111111'" (* 1425929142271 (0x14bffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~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).
-Notation "'0b11111111111111111111111111111111111011110'" (* 2199023255518 (0x1ffffffffde) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111100010'" (* 2199023255522 (0x1ffffffffe2) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1~0).
-Notation "'0b11111111111111111111111111111111111110111'" (* 2199023255543 (0x1fffffffff7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1).
-Notation "'0b11111111111111111111111111111111111111110'" (* 2199023255550 (0x1fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b11111111111111111111111111111111111111111'" (* 2199023255551 (0x1ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b100000000000000000000000000000000000000000'" (* 2199023255552 (0x20000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000001'" (* 2199023255553 (0x20000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111100000001111111111111111111111111111'" (* 4363955208191 (0x3f80fffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~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).
-Notation "'0b111111111111111111111111111111110000000000'" (* 4398046510080 (0x3fffffffc00) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b111111111111111111111111111111111111100111'" (* 4398046511079 (0x3ffffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0b111111111111111111111111111111111111101110'" (* 4398046511086 (0x3ffffffffee) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111110111'" (* 4398046511095 (0x3fffffffff7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~1~1~1).
-Notation "'0b111111111111111111111111111111111111111011'" (* 4398046511099 (0x3fffffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0b111111111111111111111111111111111111111110'" (* 4398046511102 (0x3fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b111111111111111111111111111111111111111111'" (* 4398046511103 (0x3ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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).
-Notation "'0b1000000000000000000000000000000000000000000'" (* 4398046511104 (0x40000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000001'" (* 4398046511105 (0x40000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111000000011111111111111111111111111110'" (* 8727910416382 (0x7f01ffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~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~0).
-Notation "'0b1111111111011111111111111111111110000101111'" (* 8791798053935 (0x7fefffffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b1111111111111111111110111111111111111111111'" (* 8796090925055 (0x7ffffdfffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1111111111111111111111111111111110100000011'" (* 8796093021443 (0x7fffffffd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~0~0~0~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111101000011'" (* 8796093022019 (0x7ffffffff43) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~0~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111001110'" (* 8796093022158 (0x7ffffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111100011'" (* 8796093022179 (0x7ffffffffe3) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111100111'" (* 8796093022183 (0x7ffffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111101101'" (* 8796093022189 (0x7ffffffffed) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~0~1).
-Notation "'0b1111111111111111111111111111111111111101110'" (* 8796093022190 (0x7ffffffffee) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111110001'" (* 8796093022193 (0x7fffffffff1) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111110110'" (* 8796093022198 (0x7fffffffff6) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111101'" (* 8796093022205 (0x7fffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1).
-Notation "'0b1111111111111111111111111111111111111111110'" (* 8796093022206 (0x7fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0).
-Notation "'0b1111111111111111111111111111111111111111111'" (* 8796093022207 (0x7ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~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).
-Notation "'0b10000000000000000000000000000000000000000000'" (* 8796093022208 (0x80000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000001'" (* 8796093022209 (0x80000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111110111111111111111111111100001011110'" (* 17583596107870 (0xffdfffff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0b11111111111111111111101111111111111111111110'" (* 17592181850110 (0xfffffbffffe) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111101000000110'" (* 17592186042886 (0xffffffffa06) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111010000110'" (* 17592186044038 (0xffffffffe86) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~0~0~0~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111000110'" (* 17592186044358 (0xfffffffffc6) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~0~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111001110'" (* 17592186044366 (0xfffffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111011010'" (* 17592186044378 (0xfffffffffda) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1~0~1~0).
-Notation "'0b11111111111111111111111111111111111111101111'" (* 17592186044399 (0xfffffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0b11111111111111111111111111111111111111111010'" (* 17592186044410 (0xffffffffffa) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111011'" (* 17592186044411 (0xffffffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0b11111111111111111111111111111111111111111101'" (* 17592186044413 (0xffffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1).
-Notation "'0b11111111111111111111111111111111111111111110'" (* 17592186044414 (0xffffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~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~0).
-Notation "'0b11111111111111111111111111111111111111111111'" (* 17592186044415 (0xfffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~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).
-Notation "'0b100000000000000000000000000000000000000000000'" (* 17592186044416 (0x100000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000001'" (* 17592186044417 (0x100000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111110001011'" (* 35184372088715 (0x1fffffffff8b) *)
- := (Const WO~0~0~0~0~0~0~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~0~0~0~1~0~1~1).
-Notation "'0b111111111111111111111111111111111111111000000'" (* 35184372088768 (0x1fffffffffc0) *)
- := (Const WO~0~0~0~0~0~0~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~0~0~0~0~0~0).
-Notation "'0b111111111111111111111111111111111111111011110'" (* 35184372088798 (0x1fffffffffde) *)
- := (Const WO~0~0~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111111110110'" (* 35184372088822 (0x1ffffffffff6) *)
- := (Const WO~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111010'" (* 35184372088826 (0x1ffffffffffa) *)
- := (Const WO~0~0~0~0~0~0~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~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111101'" (* 35184372088829 (0x1ffffffffffd) *)
- := (Const WO~0~0~0~0~0~0~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~0~1).
-Notation "'0b111111111111111111111111111111111111111111110'" (* 35184372088830 (0x1ffffffffffe) *)
- := (Const WO~0~0~0~0~0~0~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~0).
-Notation "'0b111111111111111111111111111111111111111111111'" (* 35184372088831 (0x1fffffffffff) *)
- := (Const WO~0~0~0~0~0~0~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).
-Notation "'0b1000000000000000000000000000000000000000000000'" (* 35184372088832 (0x200000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000001'" (* 35184372088833 (0x200000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111011111111111111111111111'" (* 70368735789055 (0x3fffff7fffff) *)
- := (Const WO~0~0~0~0~0~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~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 "'0b1111111111111111111111111111111111111100010110'" (* 70368744177430 (0x3fffffffff16) *)
- := (Const WO~0~0~0~0~0~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~0~0~0~1~0~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111100101'" (* 70368744177637 (0x3fffffffffe5) *)
- := (Const WO~0~0~0~0~0~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~0~0~1~0~1).
-Notation "'0b1111111111111111111111111111111111111111101111'" (* 70368744177647 (0x3fffffffffef) *)
- := (Const WO~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111110011'" (* 70368744177651 (0x3ffffffffff3) *)
- := (Const WO~0~0~0~0~0~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~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111110111'" (* 70368744177655 (0x3ffffffffff7) *)
- := (Const WO~0~0~0~0~0~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~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111010'" (* 70368744177658 (0x3ffffffffffa) *)
- := (Const WO~0~0~0~0~0~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~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111011'" (* 70368744177659 (0x3ffffffffffb) *)
- := (Const WO~0~0~0~0~0~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~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111100'" (* 70368744177660 (0x3ffffffffffc) *)
- := (Const WO~0~0~0~0~0~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~0~0).
-Notation "'0b1111111111111111111111111111111111111111111101'" (* 70368744177661 (0x3ffffffffffd) *)
- := (Const WO~0~0~0~0~0~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~0~1).
-Notation "'0b1111111111111111111111111111111111111111111110'" (* 70368744177662 (0x3ffffffffffe) *)
- := (Const WO~0~0~0~0~0~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~0).
-Notation "'0b1111111111111111111111111111111111111111111111'" (* 70368744177663 (0x3fffffffffff) *)
- := (Const WO~0~0~0~0~0~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).
-Notation "'0b10000000000000000000000000000000000000000000000'" (* 70368744177664 (0x400000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000001'" (* 70368744177665 (0x400000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111110111111111111111111111110'" (* 140737471578110 (0x7ffffefffffe) *)
- := (Const WO~0~0~0~0~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~0~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 "'0b11111111111111111111111011111111111111111111111'" (* 140737479966719 (0x7fffff7fffff) *)
- := (Const WO~0~0~0~0~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~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 "'0b11111111111111111111111111111111111110111000111'" (* 140737488354759 (0x7ffffffffdc7) *)
- := (Const WO~0~0~0~0~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~0~1~1~1~0~0~0~1~1~1).
-Notation "'0b11111111111111111111111111111111111111000011111'" (* 140737488354847 (0x7ffffffffe1f) *)
- := (Const WO~0~0~0~0~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~0~0~0~0~1~1~1~1~1).
-Notation "'0b11111111111111111111111111111111111111101000101'" (* 140737488355141 (0x7fffffffff45) *)
- := (Const WO~0~0~0~0~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~0~1~0~0~0~1~0~1).
-Notation "'0b11111111111111111111111111111111111111111001010'" (* 140737488355274 (0x7fffffffffca) *)
- := (Const WO~0~0~0~0~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~0~0~1~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111011110'" (* 140737488355294 (0x7fffffffffde) *)
- := (Const WO~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111100101'" (* 140737488355301 (0x7fffffffffe5) *)
- := (Const WO~0~0~0~0~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~0~0~1~0~1).
-Notation "'0b11111111111111111111111111111111111111111100110'" (* 140737488355302 (0x7fffffffffe6) *)
- := (Const WO~0~0~0~0~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~0~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111100111'" (* 140737488355303 (0x7fffffffffe7) *)
- := (Const WO~0~0~0~0~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~0~0~1~1~1).
-Notation "'0b11111111111111111111111111111111111111111101110'" (* 140737488355310 (0x7fffffffffee) *)
- := (Const WO~0~0~0~0~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~0~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111110001'" (* 140737488355313 (0x7ffffffffff1) *)
- := (Const WO~0~0~0~0~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~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111110110'" (* 140737488355318 (0x7ffffffffff6) *)
- := (Const WO~0~0~0~0~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~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111110111'" (* 140737488355319 (0x7ffffffffff7) *)
- := (Const WO~0~0~0~0~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~0~1~1~1).
-Notation "'0b11111111111111111111111111111111111111111111110'" (* 140737488355326 (0x7ffffffffffe) *)
- := (Const WO~0~0~0~0~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~0).
-Notation "'0b11111111111111111111111111111111111111111111111'" (* 140737488355327 (0x7fffffffffff) *)
- := (Const WO~0~0~0~0~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).
-Notation "'0b100000000000000000000000000000000000000000000000'" (* 140737488355328 (0x800000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000001'" (* 140737488355329 (0x800000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b101100001111111111111111111111111111111111111111'" (* 194613558116351 (0xb0ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~1~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).
-Notation "'0b111111111111101011111111111111111111111111111111'" (* 281453501874175 (0xfffaffffffff) *)
- := (Const WO~0~0~0~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~0~1~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).
-Notation "'0b111111111111111011111111111111111111111111111111'" (* 281470681743359 (0xfffeffffffff) *)
- := (Const WO~0~0~0~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~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).
-Notation "'0b111111111111111111111110111111111111111111111110'" (* 281474959933438 (0xfffffefffffe) *)
- := (Const WO~0~0~0~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~0~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 "'0b111111111111111111111111111111101111111111111111'" (* 281474976645119 (0xfffffffeffff) *)
- := (Const WO~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b111111111111111111111111111111111111101110001110'" (* 281474976709518 (0xfffffffffb8e) *)
- := (Const WO~0~0~0~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~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111110000111110'" (* 281474976709694 (0xfffffffffc3e) *)
- := (Const WO~0~0~0~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~0~0~0~0~1~1~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111111001011011'" (* 281474976710235 (0xfffffffffe5b) *)
- := (Const WO~0~0~0~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~0~0~1~0~1~1~0~1~1).
-Notation "'0b111111111111111111111111111111111111111010001010'" (* 281474976710282 (0xfffffffffe8a) *)
- := (Const WO~0~0~0~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~0~1~0~0~0~1~0~1~0).
-Notation "'0b111111111111111111111111111111111111111011000011'" (* 281474976710339 (0xfffffffffec3) *)
- := (Const WO~0~0~0~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~0~1~1~0~0~0~0~1~1).
-Notation "'0b111111111111111111111111111111111111111101000101'" (* 281474976710469 (0xffffffffff45) *)
- := (Const WO~0~0~0~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~0~1~0~0~0~1~0~1).
-Notation "'0b111111111111111111111111111111111111111110010111'" (* 281474976710551 (0xffffffffff97) *)
- := (Const WO~0~0~0~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~0~0~1~0~1~1~1).
-Notation "'0b111111111111111111111111111111111111111111001010'" (* 281474976710602 (0xffffffffffca) *)
- := (Const WO~0~0~0~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~0~0~1~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111001110'" (* 281474976710606 (0xffffffffffce) *)
- := (Const WO~0~0~0~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~0~0~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111100001'" (* 281474976710625 (0xffffffffffe1) *)
- := (Const WO~0~0~0~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~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111100010'" (* 281474976710626 (0xffffffffffe2) *)
- := (Const WO~0~0~0~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~0~0~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111100111'" (* 281474976710631 (0xffffffffffe7) *)
- := (Const WO~0~0~0~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~0~0~1~1~1).
-Notation "'0b111111111111111111111111111111111111111111101101'" (* 281474976710637 (0xffffffffffed) *)
- := (Const WO~0~0~0~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~0~1~1~0~1).
-Notation "'0b111111111111111111111111111111111111111111101110'" (* 281474976710638 (0xffffffffffee) *)
- := (Const WO~0~0~0~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~0~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111101111'" (* 281474976710639 (0xffffffffffef) *)
- := (Const WO~0~0~0~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~0~1~1~1~1).
-Notation "'0b111111111111111111111111111111111111111111110101'" (* 281474976710645 (0xfffffffffff5) *)
- := (Const WO~0~0~0~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~0~1~0~1).
-Notation "'0b111111111111111111111111111111111111111111110111'" (* 281474976710647 (0xfffffffffff7) *)
- := (Const WO~0~0~0~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~0~1~1~1).
-Notation "'0b111111111111111111111111111111111111111111111101'" (* 281474976710653 (0xfffffffffffd) *)
- := (Const WO~0~0~0~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~0~1).
-Notation "'0b111111111111111111111111111111111111111111111110'" (* 281474976710654 (0xfffffffffffe) *)
- := (Const WO~0~0~0~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~0).
-Notation "'0b111111111111111111111111111111111111111111111111'" (* 281474976710655 (0xffffffffffff) *)
- := (Const WO~0~0~0~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).
-Notation "'0b1000000000000000000000000000000000000000000000000'" (* 281474976710656 (0x1000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000001'" (* 281474976710657 (0x1000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1011000011111111111111111111111111111111111111110'" (* 389227116232702 (0x161fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~1~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~0).
-Notation "'0b1111111111111010111111111111111111111111111111110'" (* 562907003748350 (0x1fff5fffffffe) *)
- := (Const WO~0~0~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~0~1~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~0).
-Notation "'0b1111111111111111111111111111111011111111111111110'" (* 562949953290238 (0x1fffffffdfffe) *)
- := (Const WO~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111110010110110'" (* 562949953420470 (0x1fffffffffcb6) *)
- := (Const WO~0~0~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~0~0~1~0~1~1~0~1~1~0).
-Notation "'0b1111111111111111111111111111111111111110110000110'" (* 562949953420678 (0x1fffffffffd86) *)
- := (Const WO~0~0~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~0~1~1~0~0~0~0~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111010001010'" (* 562949953420938 (0x1fffffffffe8a) *)
- := (Const WO~0~0~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~0~1~0~0~0~1~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111100101110'" (* 562949953421102 (0x1ffffffffff2e) *)
- := (Const WO~0~0~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~0~0~1~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111111000010'" (* 562949953421250 (0x1ffffffffffc2) *)
- := (Const WO~0~0~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~0~0~0~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111001110'" (* 562949953421262 (0x1ffffffffffce) *)
- := (Const WO~0~0~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~0~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111111011010'" (* 562949953421274 (0x1ffffffffffda) *)
- := (Const WO~0~0~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~0~1~1~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111011110'" (* 562949953421278 (0x1ffffffffffde) *)
- := (Const WO~0~0~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~0~1~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111111011111'" (* 562949953421279 (0x1ffffffffffdf) *)
- := (Const WO~0~0~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~0~1~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111101010'" (* 562949953421290 (0x1ffffffffffea) *)
- := (Const WO~0~0~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~0~1~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111101011'" (* 562949953421291 (0x1ffffffffffeb) *)
- := (Const WO~0~0~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~0~1~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111101101'" (* 562949953421293 (0x1ffffffffffed) *)
- := (Const WO~0~0~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~0~1~1~0~1).
-Notation "'0b1111111111111111111111111111111111111111111101110'" (* 562949953421294 (0x1ffffffffffee) *)
- := (Const WO~0~0~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~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111111101111'" (* 562949953421295 (0x1ffffffffffef) *)
- := (Const WO~0~0~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~0~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111110001'" (* 562949953421297 (0x1fffffffffff1) *)
- := (Const WO~0~0~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~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111110111'" (* 562949953421303 (0x1fffffffffff7) *)
- := (Const WO~0~0~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~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111010'" (* 562949953421306 (0x1fffffffffffa) *)
- := (Const WO~0~0~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~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111100'" (* 562949953421308 (0x1fffffffffffc) *)
- := (Const WO~0~0~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~0~0).
-Notation "'0b1111111111111111111111111111111111111111111111110'" (* 562949953421310 (0x1fffffffffffe) *)
- := (Const WO~0~0~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~0).
-Notation "'0b1111111111111111111111111111111111111111111111111'" (* 562949953421311 (0x1ffffffffffff) *)
- := (Const WO~0~0~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).
-Notation "'0b10000000000000000000000000000000000000000000000000'" (* 562949953421312 (0x2000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000001'" (* 562949953421313 (0x2000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111101111111111111111111111111'" (* 1125899873288191 (0x3fffffdffffff) *)
- := (Const WO~0~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~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 "'0b11111111111111111111111111111111111111111110111110'" (* 1125899906842558 (0x3ffffffffffbe) *)
- := (Const WO~0~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~0~1~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111010110'" (* 1125899906842582 (0x3ffffffffffd6) *)
- := (Const WO~0~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~0~1~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111011010'" (* 1125899906842586 (0x3ffffffffffda) *)
- := (Const WO~0~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~0~1~1~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111011110'" (* 1125899906842590 (0x3ffffffffffde) *)
- := (Const WO~0~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~0~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111100001'" (* 1125899906842593 (0x3ffffffffffe1) *)
- := (Const WO~0~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~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111100010'" (* 1125899906842594 (0x3ffffffffffe2) *)
- := (Const WO~0~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~0~0~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111101110'" (* 1125899906842606 (0x3ffffffffffee) *)
- := (Const WO~0~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~0~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111101111'" (* 1125899906842607 (0x3ffffffffffef) *)
- := (Const WO~0~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~0~1~1~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111011'" (* 1125899906842619 (0x3fffffffffffb) *)
- := (Const WO~0~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~0~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111101'" (* 1125899906842621 (0x3fffffffffffd) *)
- := (Const WO~0~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~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111110'" (* 1125899906842622 (0x3fffffffffffe) *)
- := (Const WO~0~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~0).
-Notation "'0b11111111111111111111111111111111111111111111111111'" (* 1125899906842623 (0x3ffffffffffff) *)
- := (Const WO~0~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).
-Notation "'0b100000000000000000000000000000000000000000000000000'" (* 1125899906842624 (0x4000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000001'" (* 1125899906842625 (0x4000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b101001011111111111111111111111111111111111111111111'" (* 1460151441686527 (0x52fffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~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).
-Notation "'0b111110110111011111111111111111111111111111111111111'" (* 2211942517178367 (0x7dbbfffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~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).
-Notation "'0b111111100001010011111111111111111111111111111111111'" (* 2234929182146559 (0x7f0a7ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~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).
-Notation "'0b111111111010011111111111111111111111111111111111111'" (* 2248776156708863 (0x7fd3fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~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).
-Notation "'0b111111111111111111111111111111110000000000000000000'" (* 2251799813160960 (0x7fffffff80000) *)
- := (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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b111111111111111111111111111111111111111110100000011'" (* 2251799813684483 (0x7fffffffffd03) *)
- := (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~0~1~0~0~0~0~0~0~1~1).
-Notation "'0b111111111111111111111111111111111111111111111000010'" (* 2251799813685186 (0x7ffffffffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111111011010'" (* 2251799813685210 (0x7ffffffffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111111011110'" (* 2251799813685214 (0x7ffffffffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111100001'" (* 2251799813685217 (0x7ffffffffffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111101101'" (* 2251799813685229 (0x7ffffffffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0b111111111111111111111111111111111111111111111101111'" (* 2251799813685231 (0x7ffffffffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0b111111111111111111111111111111111111111111111110110'" (* 2251799813685238 (0x7fffffffffff6) *)
- := (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~0~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111110111'" (* 2251799813685239 (0x7fffffffffff7) *)
- := (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~0~1~1~1).
-Notation "'0b111111111111111111111111111111111111111111111111010'" (* 2251799813685242 (0x7fffffffffffa) *)
- := (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~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111101'" (* 2251799813685245 (0x7fffffffffffd) *)
- := (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~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111110'" (* 2251799813685246 (0x7fffffffffffe) *)
- := (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~0).
-Notation "'0b111111111111111111111111111111111111111111111111111'" (* 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 "'0b1000000000000000000000000000000000000000000000000000'" (* 2251799813685248 (0x8000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000001'" (* 2251799813685249 (0x8000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1010010111111111111111111111111111111111111111111110'" (* 2920302883373054 (0xa5ffffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~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~0).
-Notation "'0b1111101101110111111111111111111111111111111111111110'" (* 4423885034356734 (0xfb77ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~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~0).
-Notation "'0b1111111000010100111111111111111111111111111111111110'" (* 4469858364293118 (0xfe14ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~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~0).
-Notation "'0b1111111110100111111111111111111111111111111111111110'" (* 4497552313417726 (0xffa7ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~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~0).
-Notation "'0b1111111111111111111011111111111111111111110000101111'" (* 4503595332402223 (0xffffefffffc2f) *)
- := (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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111101000000110'" (* 4503599627368966 (0xffffffffffa06) *)
- := (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~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111110111000111'" (* 4503599627369927 (0xffffffffffdc7) *)
- := (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~0~1~1~1~0~0~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111000011111'" (* 4503599627370015 (0xffffffffffe1f) *)
- := (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~0~0~0~0~1~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111101000011'" (* 4503599627370307 (0xfffffffffff43) *)
- := (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~0~1~0~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111101000101'" (* 4503599627370309 (0xfffffffffff45) *)
- := (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~0~1~0~0~0~1~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111000010'" (* 4503599627370434 (0xfffffffffffc2) *)
- := (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~0~0~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111011010'" (* 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 "'0b1111111111111111111111111111111111111111111111011110'" (* 4503599627370462 (0xfffffffffffde) *)
- := (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~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111101011'" (* 4503599627370475 (0xfffffffffffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111101110'" (* 4503599627370478 (0xfffffffffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111101111'" (* 4503599627370479 (0xfffffffffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111010'" (* 4503599627370490 (0xffffffffffffa) *)
- := (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~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111111011'" (* 4503599627370491 (0xffffffffffffb) *)
- := (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~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111110'" (* 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).
-Notation "'0b1111111111111111111111111111111111111111111111111111'" (* 4503599627370495 (0xfffffffffffff) *)
- := (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~1).
-Notation "'0b10000000000000000000000000000000000000000000000000000'" (* 4503599627370496 (0x10000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000001'" (* 4503599627370497 (0x10000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111110111111111111111111111100001011110'" (* 9007190664804446 (0x1ffffdfffff85e) *)
- := (Const WO~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0b11111111111111111111111111011111111111111111111111111'" (* 9007199187632127 (0x1ffffffbffffff) *)
- := (Const WO~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~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 "'0b11111111111111111111111111111111111111111101110001110'" (* 9007199254739854 (0x1ffffffffffb8e) *)
- := (Const WO~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~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111110000111110'" (* 9007199254740030 (0x1ffffffffffc3e) *)
- := (Const WO~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~0~0~0~0~1~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111010000110'" (* 9007199254740614 (0x1ffffffffffe86) *)
- := (Const WO~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~0~1~0~0~0~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111010001010'" (* 9007199254740618 (0x1ffffffffffe8a) *)
- := (Const WO~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~0~1~0~0~0~1~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111010110'" (* 9007199254740950 (0x1fffffffffffd6) *)
- := (Const WO~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~0~1~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111011110'" (* 9007199254740958 (0x1fffffffffffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111100011'" (* 9007199254740963 (0x1fffffffffffe3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111100111'" (* 9007199254740967 (0x1fffffffffffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111110001'" (* 9007199254740977 (0x1ffffffffffff1) *)
- := (Const WO~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~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111110110'" (* 9007199254740982 (0x1ffffffffffff6) *)
- := (Const WO~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~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111100'" (* 9007199254740988 (0x1ffffffffffffc) *)
- := (Const WO~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~0).
-Notation "'0b11111111111111111111111111111111111111111111111111110'" (* 9007199254740990 (0x1ffffffffffffe) *)
- := (Const WO~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~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111111'" (* 9007199254740991 (0x1fffffffffffff) *)
- := (Const WO~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~1~1).
-Notation "'0b100000000000000000000000000000000000000000000000000000'" (* 9007199254740992 (0x20000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000001'" (* 9007199254740993 (0x20000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111000110'" (* 18014398509481926 (0x3fffffffffffc6) *)
- := (Const WO~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~0~0~0~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111001110'" (* 18014398509481934 (0x3fffffffffffce) *)
- := (Const WO~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~0~0~1~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111100010'" (* 18014398509481954 (0x3fffffffffffe2) *)
- := (Const WO~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~0~0~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111110111'" (* 18014398509481975 (0x3ffffffffffff7) *)
- := (Const WO~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~0~1~1~1).
-Notation "'0b111111111111111111111111111111111111111111111111111101'" (* 18014398509481981 (0x3ffffffffffffd) *)
- := (Const WO~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~1~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111110'" (* 18014398509481982 (0x3ffffffffffffe) *)
- := (Const WO~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~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111111111'" (* 18014398509481983 (0x3fffffffffffff) *)
- := (Const WO~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~1~1~1).
-Notation "'0b1000000000000000000000000000000000000000000000000000000'" (* 18014398509481984 (0x40000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000001'" (* 18014398509481985 (0x40000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111001011011'" (* 36028797018963547 (0x7ffffffffffe5b) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111011000011'" (* 36028797018963651 (0x7ffffffffffec3) *)
- := (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~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~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111101000101'" (* 36028797018963781 (0x7fffffffffff45) *)
- := (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~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~0~0~1~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111110010111'" (* 36028797018963863 (0x7fffffffffff97) *)
- := (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~1~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~0~1~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111100001'" (* 36028797018963937 (0x7fffffffffffe1) *)
- := (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~1~1~1~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~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111100111'" (* 36028797018963943 (0x7fffffffffffe7) *)
- := (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~1~1~1~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~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111101011'" (* 36028797018963947 (0x7fffffffffffeb) *)
- := (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~1~1~1~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~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111101101'" (* 36028797018963949 (0x7fffffffffffed) *)
- := (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~1~1~1~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).
-Notation "'0b1111111111111111111111111111111111111111111111111101110'" (* 36028797018963950 (0x7fffffffffffee) *)
- := (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~1~1~1~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~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111111101111'" (* 36028797018963951 (0x7fffffffffffef) *)
- := (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~1~1~1~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~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111010'" (* 36028797018963962 (0x7ffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0b1111111111111111111111111111111111111111111111111111100'" (* 36028797018963964 (0x7ffffffffffffc) *)
- := (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~1~1~1~1~1~1~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~0).
-Notation "'0b1111111111111111111111111111111111111111111111111111110'" (* 36028797018963966 (0x7ffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0b1111111111111111111111111111111111111111111111111111111'" (* 36028797018963967 (0x7fffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000'" (* 36028797018963968 (0x80000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000001'" (* 36028797018963969 (0x80000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111110000000000000000000000000000000000000000'" (* 72056494526300160 (0xffff0000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b11111111111111111111111111111111111111111111110010110110'" (* 72057594037927094 (0xfffffffffffcb6) *)
- := (Const WO~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~0~1~0~1~1~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111110110000110'" (* 72057594037927302 (0xfffffffffffd86) *)
- := (Const WO~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~0~0~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111010001010'" (* 72057594037927562 (0xfffffffffffe8a) *)
- := (Const WO~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~0~1~0~0~0~1~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111100101110'" (* 72057594037927726 (0xffffffffffff2e) *)
- := (Const WO~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~0~0~1~0~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111110001011'" (* 72057594037927819 (0xffffffffffff8b) *)
- := (Const WO~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~0~0~0~1~0~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111111000010'" (* 72057594037927874 (0xffffffffffffc2) *)
- := (Const WO~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~0~0~0~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111001110'" (* 72057594037927886 (0xffffffffffffce) *)
- := (Const WO~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~0~0~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111010110'" (* 72057594037927894 (0xffffffffffffd6) *)
- := (Const WO~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~0~1~0~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111011010'" (* 72057594037927898 (0xffffffffffffda) *)
- := (Const WO~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~0~1~1~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111011110'" (* 72057594037927902 (0xffffffffffffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111101011'" (* 72057594037927915 (0xffffffffffffeb) *)
- := (Const WO~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~1~0~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111111101111'" (* 72057594037927919 (0xffffffffffffef) *)
- := (Const WO~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~1~1~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111011'" (* 72057594037927931 (0xfffffffffffffb) *)
- := (Const WO~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~1~1~0~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111101'" (* 72057594037927933 (0xfffffffffffffd) *)
- := (Const WO~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~1~1~1~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111110'" (* 72057594037927934 (0xfffffffffffffe) *)
- := (Const WO~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~1~1~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111111111'" (* 72057594037927935 (0xffffffffffffff) *)
- := (Const WO~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~1~1~1~1~1).
-Notation "'0b100000000000000000000000000000000000000000000000000000000'" (* 72057594037927936 (0x100000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000001'" (* 72057594037927937 (0x100000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111100010110'" (* 144115188075855638 (0x1ffffffffffff16) *)
- := (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~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~0~0~1~0~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111111010110'" (* 144115188075855830 (0x1ffffffffffffd6) *)
- := (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~1~1~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~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111111011110'" (* 144115188075855838 (0x1ffffffffffffde) *)
- := (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~1~1~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~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111111101101'" (* 144115188075855853 (0x1ffffffffffffed) *)
- := (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~1~1~1~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).
-Notation "'0b111111111111111111111111111111111111111111111111111110001'" (* 144115188075855857 (0x1fffffffffffff1) *)
- := (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~1~1~1~1~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~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111110110'" (* 144115188075855862 (0x1fffffffffffff6) *)
- := (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~1~1~1~1~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).
-Notation "'0b111111111111111111111111111111111111111111111111111110111'" (* 144115188075855863 (0x1fffffffffffff7) *)
- := (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~1~1~1~1~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~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111010'" (* 144115188075855866 (0x1fffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0b111111111111111111111111111111111111111111111111111111011'" (* 144115188075855867 (0x1fffffffffffffb) *)
- := (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~1~1~1~1~1~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).
-Notation "'0b111111111111111111111111111111111111111111111111111111100'" (* 144115188075855868 (0x1fffffffffffffc) *)
- := (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~1~1~1~1~1~1~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~0).
-Notation "'0b111111111111111111111111111111111111111111111111111111101'" (* 144115188075855869 (0x1fffffffffffffd) *)
- := (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~1~1~1~1~1~1~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).
-Notation "'0b111111111111111111111111111111111111111111111111111111110'" (* 144115188075855870 (0x1fffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0b111111111111111111111111111111111111111111111111111111111'" (* 144115188075855871 (0x1ffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000'" (* 144115188075855872 (0x200000000000000) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000001'" (* 144115188075855873 (0x200000000000001) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111011010'" (* 288230376151711706 (0x3ffffffffffffda) *)
- := (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~1~1~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 "'0b1111111111111111111111111111111111111111111111111111100001'" (* 288230376151711713 (0x3ffffffffffffe1) *)
- := (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~1~1~1~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~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111100010'" (* 288230376151711714 (0x3ffffffffffffe2) *)
- := (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~1~1~1~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~0~0~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111111111100101'" (* 288230376151711717 (0x3ffffffffffffe5) *)
- := (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~1~1~1~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~0~1~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111101110'" (* 288230376151711726 (0x3ffffffffffffee) *)
- := (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~1~1~1~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~1~0).
-Notation "'0b1111111111111111111111111111111111111111111111111111101111'" (* 288230376151711727 (0x3ffffffffffffef) *)
- := (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~1~1~1~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~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111110110'" (* 288230376151711734 (0x3fffffffffffff6) *)
- := (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~1~1~1~1~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).
-Notation "'0b1111111111111111111111111111111111111111111111111111111010'" (* 288230376151711738 (0x3fffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0b1111111111111111111111111111111111111111111111111111111100'" (* 288230376151711740 (0x3fffffffffffffc) *)
- := (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~1~1~1~1~1~1~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~0).
-Notation "'0b1111111111111111111111111111111111111111111111111111111101'" (* 288230376151711741 (0x3fffffffffffffd) *)
- := (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~1~1~1~1~1~1~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).
-Notation "'0b1111111111111111111111111111111111111111111111111111111110'" (* 288230376151711742 (0x3fffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0b1111111111111111111111111111111111111111111111111111111111'" (* 288230376151711743 (0x3ffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000'" (* 288230376151711744 (0x400000000000000) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000001'" (* 288230376151711745 (0x400000000000001) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111000010'" (* 576460752303423426 (0x7ffffffffffffc2) *)
- := (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~1~1~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~0~0~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111111001010'" (* 576460752303423434 (0x7ffffffffffffca) *)
- := (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~1~1~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~0~1~0~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111111011110'" (* 576460752303423454 (0x7ffffffffffffde) *)
- := (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~1~1~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~1~1~0).
-Notation "'0b11111111111111111111111111111111111111111111111111111101011'" (* 576460752303423467 (0x7ffffffffffffeb) *)
- := (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~1~1~1~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~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111101101'" (* 576460752303423469 (0x7ffffffffffffed) *)
- := (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~1~1~1~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).
-Notation "'0b11111111111111111111111111111111111111111111111111111101111'" (* 576460752303423471 (0x7ffffffffffffef) *)
- := (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~1~1~1~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~1~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111110001'" (* 576460752303423473 (0x7fffffffffffff1) *)
- := (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~1~1~1~1~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~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111010'" (* 576460752303423482 (0x7fffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0b11111111111111111111111111111111111111111111111111111111110'" (* 576460752303423486 (0x7fffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0b11111111111111111111111111111111111111111111111111111111111'" (* 576460752303423487 (0x7ffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000'" (* 576460752303423488 (0x800000000000000) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000001'" (* 576460752303423489 (0x800000000000001) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111110000011111000001111100000111110000011111000001111100001'" (* 1117984489315730401 (0xf83e0f83e0f83e1) *)
- := (Const WO~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111010110'" (* 1152921504606846934 (0xfffffffffffffd6) *)
- := (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~1~1~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~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111111111011010'" (* 1152921504606846938 (0xfffffffffffffda) *)
- := (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~1~1~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 "'0b111111111111111111111111111111111111111111111111111111011110'" (* 1152921504606846942 (0xfffffffffffffde) *)
- := (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~1~1~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~1~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111111111100010'" (* 1152921504606846946 (0xfffffffffffffe2) *)
- := (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~1~1~1~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~0~0~1~0).
-Notation "'0b111111111111111111111111111111111111111111111111111111111110'" (* 1152921504606846974 (0xffffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0b111111111111111111111111111111111111111111111111111111111111'" (* 1152921504606846975 (0xfffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000'" (* 1152921504606846976 (0x1000000000000000) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000001'" (* 1152921504606846977 (0x1000000000000001) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111100'" (* 2305843009213693948 (0x1ffffffffffffffc) *)
- := (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~1~1~1~1~1~1~1~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~0).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111110'" (* 2305843009213693950 (0x1ffffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0b1111111111111111111111111111111111111111111111111111111111111'" (* 2305843009213693951 (0x1fffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000'" (* 2305843009213693952 (0x2000000000000000) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000001'" (* 2305843009213693953 (0x2000000000000001) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b10111010001011101000101110100010111010001011101000101110100011'" (* 3353953467947191203 (0x2e8ba2e8ba2e8ba3) *)
- := (Const WO~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~1).
-Notation "'0b11010011110111001011000010001101001111011100101100001000110101'" (* 3816567739388183093 (0x34f72c234f72c235) *)
- := (Const WO~0~0~1~1~0~1~0~0~1~1~1~1~0~1~1~1~0~0~1~0~1~1~0~0~0~0~1~0~0~0~1~1~0~1~0~0~1~1~1~1~0~1~1~1~0~0~1~0~1~1~0~0~0~0~1~0~0~0~1~1~0~1~0~1).
-Notation "'0b11111011011101111111111111111111111111111111111111111111111111'" (* 4530058275181297663 (0x3eddffffffffffff) *)
- := (Const WO~0~0~1~1~1~1~1~0~1~1~0~1~1~1~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).
-Notation "'0b11111110000000111111111111111111111111111111111111111111111111'" (* 4575938696385134591 (0x3f80ffffffffffff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~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).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111'" (* 4611686018427387903 (0x3fffffffffffffff) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000'" (* 4611686018427387904 (0x4000000000000000) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000001'" (* 4611686018427387905 (0x4000000000000001) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b100111011000100111011000100111011000100111011000100111011000101'" (* 5675921253449092805 (0x4ec4ec4ec4ec4ec5) *)
- := (Const WO~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111'" (* 9223372036854775807 (0x7fffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000'" (* 9223372036854775808 (0x8000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000001'" (* 9223372036854775809 (0x8000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1000010010111101101000010010111101101000010010111101101000010011'" (* 9564978408590137875 (0x84bda12f684bda13L) *)
- := (Const WO~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~1).
-Notation "'0b1000011010111100101000011010111100101000011010111100101000011011'" (* 9708812670373448219 (0x86bca1af286bca1bL) *)
- := (Const WO~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~1).
-Notation "'0b1000101001000100011100101111111010100001100010100100010001110011'" (* 9963214713607832691 (0x8a4472fea18a4473L) *)
- := (Const WO~1~0~0~0~1~0~1~0~0~1~0~0~0~1~0~0~0~1~1~1~0~0~1~0~1~1~1~1~1~1~1~0~1~0~1~0~0~0~0~1~1~0~0~0~1~0~1~0~0~1~0~0~0~1~0~0~0~1~1~1~0~0~1~1).
-Notation "'0b1000111000111000111000111000111000111000111000111000111000111001'" (* 10248191152060862009 (0x8e38e38e38e38e39L) *)
- := (Const WO~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~1).
-Notation "'0b1000111101011100001010001111010111000010100011110101110000101001'" (* 10330176681277348905 (0x8f5c28f5c28f5c29L) *)
- := (Const WO~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~1).
-Notation "'0b1000111111011000111111011000111111011000111111011000111111011001'" (* 10365313336655843289 (0x8fd8fd8fd8fd8fd9L) *)
- := (Const WO~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~1).
-Notation "'0b1010010100111111101010010100111111101010010100111111101010010101'" (* 11907422100489763477 (0xa53fa94fea53fa95L) *)
- := (Const WO~1~0~1~0~0~1~0~1~0~0~1~1~1~1~1~1~1~0~1~0~1~0~0~1~0~1~0~0~1~1~1~1~1~1~1~0~1~0~1~0~0~1~0~1~0~0~1~1~1~1~1~1~1~0~1~0~1~0~0~1~0~1~0~1).
-Notation "'0b1010101001010100111111111010101001010100111111111010101001010101'" (* 12273715991527008853 (0xaa54ffaa54ffaa55L) *)
- := (Const WO~1~0~1~0~1~0~1~0~0~1~0~1~0~1~0~0~1~1~1~1~1~1~1~1~1~0~1~0~1~0~1~0~0~1~0~1~0~1~0~0~1~1~1~1~1~1~1~1~1~0~1~0~1~0~1~0~0~1~0~1~0~1~0~1).
-Notation "'0b1010101010101010101010101010101010101010101010101010101010101011'" (* 12297829382473034411 (0xaaaaaaaaaaaaaaabL) *)
- := (Const WO~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~1).
-Notation "'0b1011000011111111111111111111111111111111111111111111111111111111'" (* 12754194144713244671 (0xb0ffffffffffffffL) *)
- := (Const WO~1~0~1~1~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~1~1~1~1~1).
-Notation "'0b1100110011001100110011001100110011001100110011001100110011001101'" (* 14757395258967641293 (0xcccccccccccccccdL) *)
- := (Const WO~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~1).
-Notation "'0b1100111010111110111011111001010011111010100001101111111000101101'" (* 14897608040525528621 (0xcebeef94fa86fe2dL) *)
- := (Const WO~1~1~0~0~1~1~1~0~1~0~1~1~1~1~1~0~1~1~1~0~1~1~1~1~1~0~0~1~0~1~0~0~1~1~1~1~1~0~1~0~1~0~0~0~0~1~1~0~1~1~1~1~1~1~1~0~0~0~1~0~1~1~0~1).
-Notation "'0b1100111100111100111100111100111100111100111100111100111100111101'" (* 14933078535860113213 (0xcf3cf3cf3cf3cf3dL) *)
- := (Const WO~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~1).
-Notation "'0b1100111111011100111111011100111111011100111111011100111111011101'" (* 14978125529935106013 (0xcfdcfdcfdcfdcfddL) *)
- := (Const WO~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~1).
-Notation "'0b1101100000111000000010010001110111010010001001010011010100110001'" (* 15580212934572586289 (0xd838091dd2253531L) *)
- := (Const WO~1~1~0~1~1~0~0~0~0~0~1~1~1~0~0~0~0~0~0~0~1~0~0~1~0~0~0~1~1~1~0~1~1~1~0~1~0~0~1~0~0~0~1~0~0~1~0~1~0~0~1~1~0~1~0~1~0~0~1~1~0~0~0~1).
-Notation "'0b1110110010011110010010001010111001101111011100011101111000010101'" (* 17050145153302519317 (0xec9e48ae6f71de15L) *)
- := (Const WO~1~1~1~0~1~1~0~0~1~0~0~1~1~1~1~0~0~1~0~0~1~0~0~0~1~0~1~0~1~1~1~0~0~1~1~0~1~1~1~1~0~1~1~1~0~0~0~1~1~1~0~1~1~1~1~0~0~0~0~1~0~1~0~1).
-Notation "'0b1110111011101110111011101110111011101110111011101110111011101111'" (* 17216961135462248175 (0xeeeeeeeeeeeeeeefL) *)
- := (Const WO~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~1).
-Notation "'0b1110111101111011110111101111011110111101111011110111101111011111'" (* 17256631552825064415 (0xef7bdef7bdef7bdfL) *)
- := (Const WO~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~1).
-Notation "'0b1111000011110000111100001111000011110000111100001111000011110001'" (* 17361641481138401521 (0xf0f0f0f0f0f0f0f1L) *)
- := (Const WO~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~1).
-Notation "'0b1111111000010100111111111111111111111111111111111111111111111111'" (* 18308539860144619519 (0xfe14ffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~0~0~0~0~1~0~1~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).
-Notation "'0b1111111110100111111111111111111111111111111111111111111111111111'" (* 18421974275759013887 (0xffa7ffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~0~1~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 "'0b1111111111111010111111111111111111111111111111111111111111111111'" (* 18445336698825998335 (0xfffaffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~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).
-Notation "'0b1111111111111110111111111111111111111111111111111111111111111111'" (* 18446462598732840959 (0xfffeffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0b1111111111111111111011111111111111111111111111111111111111111111'" (* 18446726481523507199 (0xffffefffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0b1111111111111111111111111111110111111111111111111111111111111111'" (* 18446744065119617023 (0xfffffffdffffffffL) *)
- := (Const WO~1~1~1~1~1~1~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~1~1~1~1~1~1~1~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 "'0b1111111111111111111111111111111011111111111111111111110000101111'" (* 18446744069414583343 (0xfffffffefffffc2fL) *)
- := (Const WO~1~1~1~1~1~1~1~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0b1111111111111111111111111111111011111111111111111111111111111111'" (* 18446744069414584319 (0xfffffffeffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~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~1~1~1~1~1~1~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 "'0b1111111111111111111111111111111100000000000000000000000000000000'" (* 18446744069414584320 (0xffffffff00000000L) *)
- := (Const WO~1~1~1~1~1~1~1~1~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1111111111111111111111111111111100000000000000000000000000000001'" (* 18446744069414584321 (0xffffffff00000001L) *)
- := (Const WO~1~1~1~1~1~1~1~1~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111101111111111111111'" (* 18446744073709486079 (0xfffffffffffeffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111110100000011'" (* 18446744073709550851 (0xfffffffffffffd03L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111110111000111'" (* 18446744073709551047 (0xfffffffffffffdc7L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~0~0~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111000011111'" (* 18446744073709551135 (0xfffffffffffffe1fL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1~1~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111001011011'" (* 18446744073709551195 (0xfffffffffffffe5bL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~0~1~1~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111011000011'" (* 18446744073709551299 (0xfffffffffffffec3L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111101000011'" (* 18446744073709551427 (0xffffffffffffff43L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111101000101'" (* 18446744073709551429 (0xffffffffffffff45L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111110001011'" (* 18446744073709551499 (0xffffffffffffff8bL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111110010111'" (* 18446744073709551511 (0xffffffffffffff97L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111011111'" (* 18446744073709551583 (0xffffffffffffffdfL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111100001'" (* 18446744073709551585 (0xffffffffffffffe1L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111100011'" (* 18446744073709551587 (0xffffffffffffffe3L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111100101'" (* 18446744073709551589 (0xffffffffffffffe5L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111100111'" (* 18446744073709551591 (0xffffffffffffffe7L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111101011'" (* 18446744073709551595 (0xffffffffffffffebL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111101101'" (* 18446744073709551597 (0xffffffffffffffedL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111101111'" (* 18446744073709551599 (0xffffffffffffffefL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111110001'" (* 18446744073709551601 (0xfffffffffffffff1L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111110011'" (* 18446744073709551603 (0xfffffffffffffff3L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111110101'" (* 18446744073709551605 (0xfffffffffffffff5L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111110111'" (* 18446744073709551607 (0xfffffffffffffff7L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111011'" (* 18446744073709551611 (0xfffffffffffffffbL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111101'" (* 18446744073709551613 (0xfffffffffffffffdL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111110'" (* 18446744073709551614 (0xfffffffffffffffeL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1111111111111111111111111111111111111111111111111111111111111111'" (* 18446744073709551615 (0xffffffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000'" (* 18446744073709551616 (0x10000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000001'" (* 18446744073709551617 (0x10000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111'" (* 36893488147419103231 (0x1ffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000'" (* 36893488147419103232 (0x20000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000001'" (* 36893488147419103233 (0x20000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111'" (* 73786976294838206463 (0x3ffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000'" (* 73786976294838206464 (0x40000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000001'" (* 73786976294838206465 (0x40000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111'" (* 147573952589676412927 (0x7ffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000'" (* 147573952589676412928 (0x80000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000001'" (* 147573952589676412929 (0x80000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111'" (* 295147905179352825855 (0xfffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000'" (* 295147905179352825856 (0x100000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000001'" (* 295147905179352825857 (0x100000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111'" (* 590295810358705651711 (0x1fffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000'" (* 590295810358705651712 (0x200000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000001'" (* 590295810358705651713 (0x200000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111'" (* 1180591620717411303423 (0x3fffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000'" (* 1180591620717411303424 (0x400000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000001'" (* 1180591620717411303425 (0x400000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111'" (* 2361183241434822606847 (0x7fffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000'" (* 2361183241434822606848 (0x800000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000001'" (* 2361183241434822606849 (0x800000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111'" (* 4722366482869645213695 (0xffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000'" (* 4722366482869645213696 (0x1000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000001'" (* 4722366482869645213697 (0x1000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111'" (* 9444732965739290427391 (0x1ffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000'" (* 9444732965739290427392 (0x2000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000001'" (* 9444732965739290427393 (0x2000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111'" (* 18889465931478580854783 (0x3ffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000'" (* 18889465931478580854784 (0x4000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000001'" (* 18889465931478580854785 (0x4000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 37778931862957161709567 (0x7ffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 37778931862957161709568 (0x8000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 37778931862957161709569 (0x8000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 75557863725914323419135 (0xfffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 75557863725914323419136 (0x10000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 75557863725914323419137 (0x10000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 151115727451828646838271 (0x1fffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 151115727451828646838272 (0x20000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 151115727451828646838273 (0x20000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 302231454903657293676543 (0x3fffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 302231454903657293676544 (0x40000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 302231454903657293676545 (0x40000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 604462909807314587353087 (0x7fffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 604462909807314587353088 (0x80000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 604462909807314587353089 (0x80000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1208925819614629174706175 (0xffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1208925819614629174706176 (0x100000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1208925819614629174706177 (0x100000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2417851639229258349412351 (0x1ffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2417851639229258349412352 (0x200000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2417851639229258349412353 (0x200000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 4835703278458516698824703 (0x3ffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 4835703278458516698824704 (0x400000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 4835703278458516698824705 (0x400000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 9671406556917033397649407 (0x7ffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 9671406556917033397649408 (0x800000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 9671406556917033397649409 (0x800000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 19342813113834066795298815 (0xfffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 19342813113834066795298816 (0x1000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 19342813113834066795298817 (0x1000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 38685626227668133590597631 (0x1fffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 38685626227668133590597632 (0x2000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 38685626227668133590597633 (0x2000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111110110'" (* 77371252455336267181195254 (0x3ffffffffffffffffffff6L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~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).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111110'" (* 77371252455336267181195262 (0x3ffffffffffffffffffffeL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~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 "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 77371252455336267181195263 (0x3fffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 77371252455336267181195264 (0x4000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 77371252455336267181195265 (0x4000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 154742504910672534362390527 (0x7fffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 154742504910672534362390528 (0x8000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 154742504910672534362390529 (0x8000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 309485009821345068724781055 (0xffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 309485009821345068724781056 (0x10000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 309485009821345068724781057 (0x10000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 618970019642690137449562111 (0x1ffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 618970019642690137449562112 (0x20000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 618970019642690137449562113 (0x20000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1237940039285380274899124223 (0x3ffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1237940039285380274899124224 (0x40000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1237940039285380274899124225 (0x40000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2475880078570760549798248447 (0x7ffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2475880078570760549798248448 (0x80000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2475880078570760549798248449 (0x80000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 4951760157141521099596496895 (0xfffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 4951760157141521099596496896 (0x100000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 4951760157141521099596496897 (0x100000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 9903520314283042199192993791 (0x1fffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 9903520314283042199192993792 (0x200000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 9903520314283042199192993793 (0x200000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 19807040628566084398385987583 (0x3fffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 19807040628566084398385987584 (0x400000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 19807040628566084398385987585 (0x400000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 39614081257132168796771975167 (0x7fffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 39614081257132168796771975168 (0x800000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 39614081257132168796771975169 (0x800000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 79228162514264337593543950335 (0xffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 79228162514264337593543950336 (0x1000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 79228162514264337593543950337 (0x1000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 158456325028528675187087900671 (0x1ffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 158456325028528675187087900672 (0x2000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 158456325028528675187087900673 (0x2000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 316912650057057350374175801343 (0x3ffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 316912650057057350374175801344 (0x4000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 316912650057057350374175801345 (0x4000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 633825300114114700748351602687 (0x7ffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 633825300114114700748351602688 (0x8000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 633825300114114700748351602689 (0x8000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1267650600228229401496703205375 (0xfffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1267650600228229401496703205376 (0x10000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1267650600228229401496703205377 (0x10000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2535301200456458802993406410751 (0x1fffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2535301200456458802993406410752 (0x20000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2535301200456458802993406410753 (0x20000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5070602400912917605986812821503 (0x3fffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5070602400912917605986812821504 (0x40000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5070602400912917605986812821505 (0x40000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 10141204801825835211973625643007 (0x7fffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 10141204801825835211973625643008 (0x80000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 10141204801825835211973625643009 (0x80000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 20282409603651670423947251286015 (0xffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 20282409603651670423947251286016 (0x100000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 20282409603651670423947251286017 (0x100000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 40564819207303340847894502572031 (0x1ffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 40564819207303340847894502572032 (0x200000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 40564819207303340847894502572033 (0x200000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 81129638414606681695789005144063 (0x3ffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 81129638414606681695789005144064 (0x400000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 81129638414606681695789005144065 (0x400000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 162259276829213363391578010288127 (0x7ffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 162259276829213363391578010288128 (0x800000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 162259276829213363391578010288129 (0x800000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 324518553658426726783156020576255 (0xfffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 324518553658426726783156020576256 (0x1000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 324518553658426726783156020576257 (0x1000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 649037107316853453566312041152511 (0x1fffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 649037107316853453566312041152512 (0x2000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 649037107316853453566312041152513 (0x2000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1298074214633706907132624082305023 (0x3fffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1298074214633706907132624082305024 (0x4000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1298074214633706907132624082305025 (0x4000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2596148429267413814265248164610047 (0x7fffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2596148429267413814265248164610048 (0x8000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2596148429267413814265248164610049 (0x8000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5192296858534827628530496329220095 (0xffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5192296858534827628530496329220096 (0x10000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5192296858534827628530496329220097 (0x10000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 10384593717069655257060992658440191 (0x1ffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 10384593717069655257060992658440192 (0x20000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 10384593717069655257060992658440193 (0x20000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 20769187434139310514121985316880383 (0x3ffffffffffffffffffffffffffffL) *)
- := (Const WO~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 20769187434139310514121985316880384 (0x40000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 20769187434139310514121985316880385 (0x40000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 41538374868278621028243970633760767 (0x7ffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 41538374868278621028243970633760768 (0x80000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 41538374868278621028243970633760769 (0x80000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 83076749736557242056487941267521535 (0xfffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 83076749736557242056487941267521536 (0x100000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 83076749736557242056487941267521537 (0x100000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 166153499473114484112975882535043071 (0x1fffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 166153499473114484112975882535043072 (0x200000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 166153499473114484112975882535043073 (0x200000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 332306998946228968225951765070086143 (0x3fffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 332306998946228968225951765070086144 (0x400000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 332306998946228968225951765070086145 (0x400000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 664613997892457936451903530140172287 (0x7fffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 664613997892457936451903530140172288 (0x800000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 664613997892457936451903530140172289 (0x800000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1329227995784915872903807060280344575 (0xffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1329227995784915872903807060280344576 (0x1000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1329227995784915872903807060280344577 (0x1000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2658455991569831745807614120560689151 (0x1ffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2658455991569831745807614120560689152 (0x2000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2658455991569831745807614120560689153 (0x2000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5316911983139663491615228241121378303 (0x3ffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5316911983139663491615228241121378304 (0x4000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5316911983139663491615228241121378305 (0x4000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 10633823966279326983230456482242756607 (0x7ffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 10633823966279326983230456482242756608 (0x8000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 10633823966279326983230456482242756609 (0x8000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 21267647932558653966460912964485513215 (0xfffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 21267647932558653966460912964485513216 (0x10000000000000000000000000000000L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 21267647932558653966460912964485513217 (0x10000000000000000000000000000001L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 42535295865117307932921825928971026431 (0x1fffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 42535295865117307932921825928971026432 (0x20000000000000000000000000000000L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 42535295865117307932921825928971026433 (0x20000000000000000000000000000001L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 85070591730234615865843651857942052863 (0x3fffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 85070591730234615865843651857942052864 (0x40000000000000000000000000000000L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 85070591730234615865843651857942052865 (0x40000000000000000000000000000001L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 170141183460469231731687303715884105727 (0x7fffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 170141183460469231731687303715884105728 (0x80000000000000000000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 170141183460469231731687303715884105729 (0x80000000000000000000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111000000000000000000000000000000010000000000000000000000000000000000000000000000000000000000000000'" (* 340282366841710300967557013911933812736 (0xffffffff000000010000000000000000L) *)
- := (Const WO~1~1~1~1~1~1~1~1~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 340282366920938463463374607431768211455 (0xffffffffffffffffffffffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 340282366920938463463374607431768211456 (0x100000000000000000000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 340282366920938463463374607431768211457 (0x100000000000000000000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 680564733841876926926749214863536422911 (0x1ffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 680564733841876926926749214863536422912 (0x200000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 680564733841876926926749214863536422913 (0x200000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1361129467683753853853498429727072845823 (0x3ffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1361129467683753853853498429727072845824 (0x400000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1361129467683753853853498429727072845825 (0x400000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2722258935367507707706996859454145691647 (0x7ffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2722258935367507707706996859454145691648 (0x800000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2722258935367507707706996859454145691649 (0x800000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5444517870735015415413993718908291383295 (0xfffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5444517870735015415413993718908291383296 (0x1000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5444517870735015415413993718908291383297 (0x1000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 10889035741470030830827987437816582766591 (0x1fffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 10889035741470030830827987437816582766592 (0x2000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 10889035741470030830827987437816582766593 (0x2000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 21778071482940061661655974875633165533183 (0x3fffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 21778071482940061661655974875633165533184 (0x4000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 21778071482940061661655974875633165533185 (0x4000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 43556142965880123323311949751266331066367 (0x7fffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 43556142965880123323311949751266331066368 (0x8000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 43556142965880123323311949751266331066369 (0x8000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 87112285931760246646623899502532662132735 (0xffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 87112285931760246646623899502532662132736 (0x10000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 87112285931760246646623899502532662132737 (0x10000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 174224571863520493293247799005065324265471 (0x1ffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 174224571863520493293247799005065324265472 (0x20000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 174224571863520493293247799005065324265473 (0x20000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 348449143727040986586495598010130648530943 (0x3ffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 348449143727040986586495598010130648530944 (0x40000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 348449143727040986586495598010130648530945 (0x40000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 696898287454081973172991196020261297061887 (0x7ffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 696898287454081973172991196020261297061888 (0x80000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 696898287454081973172991196020261297061889 (0x80000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1393796574908163946345982392040522594123775 (0xfffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1393796574908163946345982392040522594123776 (0x100000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1393796574908163946345982392040522594123777 (0x100000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2787593149816327892691964784081045188247551 (0x1fffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2787593149816327892691964784081045188247552 (0x200000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2787593149816327892691964784081045188247553 (0x200000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5575186299632655785383929568162090376495103 (0x3fffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5575186299632655785383929568162090376495104 (0x400000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5575186299632655785383929568162090376495105 (0x400000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 11150372599265311570767859136324180752990207 (0x7fffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 11150372599265311570767859136324180752990208 (0x800000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 11150372599265311570767859136324180752990209 (0x800000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 22300745198530623141535718272648361505980415 (0xffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 22300745198530623141535718272648361505980416 (0x1000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 22300745198530623141535718272648361505980417 (0x1000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 44601490397061246283071436545296723011960831 (0x1ffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 44601490397061246283071436545296723011960832 (0x2000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 44601490397061246283071436545296723011960833 (0x2000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 89202980794122492566142873090593446023921663 (0x3ffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 89202980794122492566142873090593446023921664 (0x4000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 89202980794122492566142873090593446023921665 (0x4000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 178405961588244985132285746181186892047843327 (0x7ffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 178405961588244985132285746181186892047843328 (0x8000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 178405961588244985132285746181186892047843329 (0x8000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 356811923176489970264571492362373784095686655 (0xfffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 356811923176489970264571492362373784095686656 (0x10000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 356811923176489970264571492362373784095686657 (0x10000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 713623846352979940529142984724747568191373311 (0x1fffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 713623846352979940529142984724747568191373312 (0x20000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 713623846352979940529142984724747568191373313 (0x20000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1427247692705959881058285969449495136382746623 (0x3fffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1427247692705959881058285969449495136382746624 (0x40000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1427247692705959881058285969449495136382746625 (0x40000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2854495385411919762116571938898990272765493247 (0x7fffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2854495385411919762116571938898990272765493248 (0x80000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2854495385411919762116571938898990272765493249 (0x80000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5708990770823839524233143877797980545530986495 (0xffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5708990770823839524233143877797980545530986496 (0x100000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5708990770823839524233143877797980545530986497 (0x100000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 11417981541647679048466287755595961091061972991 (0x1ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 11417981541647679048466287755595961091061972992 (0x200000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 11417981541647679048466287755595961091061972993 (0x200000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 22835963083295358096932575511191922182123945983 (0x3ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 22835963083295358096932575511191922182123945984 (0x400000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 22835963083295358096932575511191922182123945985 (0x400000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 45671926166590716193865151022383844364247891967 (0x7ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 45671926166590716193865151022383844364247891968 (0x800000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 45671926166590716193865151022383844364247891969 (0x800000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 91343852333181432387730302044767688728495783935 (0xfffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 91343852333181432387730302044767688728495783936 (0x1000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 91343852333181432387730302044767688728495783937 (0x1000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 182687704666362864775460604089535377456991567871 (0x1fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 182687704666362864775460604089535377456991567872 (0x2000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 182687704666362864775460604089535377456991567873 (0x2000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 365375409332725729550921208179070754913983135743 (0x3fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 365375409332725729550921208179070754913983135744 (0x4000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 365375409332725729550921208179070754913983135745 (0x4000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 730750818665451459101842416358141509827966271487 (0x7fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 730750818665451459101842416358141509827966271488 (0x8000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 730750818665451459101842416358141509827966271489 (0x8000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1461501637330902918203684832716283019655932542975 (0xffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1461501637330902918203684832716283019655932542976 (0x10000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1461501637330902918203684832716283019655932542977 (0x10000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2923003274661805836407369665432566039311865085951 (0x1ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2923003274661805836407369665432566039311865085952 (0x20000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2923003274661805836407369665432566039311865085953 (0x20000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5846006549323611672814739330865132078623730171903 (0x3ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5846006549323611672814739330865132078623730171904 (0x40000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5846006549323611672814739330865132078623730171905 (0x40000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 11692013098647223345629478661730264157247460343807 (0x7ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 11692013098647223345629478661730264157247460343808 (0x80000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 11692013098647223345629478661730264157247460343809 (0x80000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 23384026197294446691258957323460528314494920687615 (0xfffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 23384026197294446691258957323460528314494920687616 (0x100000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 23384026197294446691258957323460528314494920687617 (0x100000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 46768052394588893382517914646921056628989841375231 (0x1fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 46768052394588893382517914646921056628989841375232 (0x200000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 46768052394588893382517914646921056628989841375233 (0x200000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 93536104789177786765035829293842113257979682750463 (0x3fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 93536104789177786765035829293842113257979682750464 (0x400000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 93536104789177786765035829293842113257979682750465 (0x400000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 187072209578355573530071658587684226515959365500927 (0x7fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 187072209578355573530071658587684226515959365500928 (0x800000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 187072209578355573530071658587684226515959365500929 (0x800000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 374144419156711147060143317175368453031918731001855 (0xffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 374144419156711147060143317175368453031918731001856 (0x1000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 374144419156711147060143317175368453031918731001857 (0x1000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 748288838313422294120286634350736906063837462003711 (0x1ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 748288838313422294120286634350736906063837462003712 (0x2000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 748288838313422294120286634350736906063837462003713 (0x2000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1496577676626844588240573268701473812127674924007423 (0x3ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1496577676626844588240573268701473812127674924007424 (0x4000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1496577676626844588240573268701473812127674924007425 (0x4000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 2993155353253689176481146537402947624255349848014847 (0x7ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 2993155353253689176481146537402947624255349848014848 (0x8000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 2993155353253689176481146537402947624255349848014849 (0x8000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 5986310706507378352962293074805895248510699696029695 (0xfffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 5986310706507378352962293074805895248510699696029696 (0x10000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 5986310706507378352962293074805895248510699696029697 (0x10000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 11972621413014756705924586149611790497021399392059391 (0x1fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 11972621413014756705924586149611790497021399392059392 (0x20000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 11972621413014756705924586149611790497021399392059393 (0x20000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 23945242826029513411849172299223580994042798784118783 (0x3fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 23945242826029513411849172299223580994042798784118784 (0x40000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 23945242826029513411849172299223580994042798784118785 (0x40000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 47890485652059026823698344598447161988085597568237567 (0x7fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 47890485652059026823698344598447161988085597568237568 (0x80000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 47890485652059026823698344598447161988085597568237569 (0x80000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 95780971304118053647396689196894323976171195136475135 (0xffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 95780971304118053647396689196894323976171195136475136 (0x100000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 95780971304118053647396689196894323976171195136475137 (0x100000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 191561942608236107294793378393788647952342390272950271 (0x1ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 191561942608236107294793378393788647952342390272950272 (0x200000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 191561942608236107294793378393788647952342390272950273 (0x200000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 383123885216472214589586756787577295904684780545900543 (0x3ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 383123885216472214589586756787577295904684780545900544 (0x400000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 383123885216472214589586756787577295904684780545900545 (0x400000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 766247770432944429179173513575154591809369561091801087 (0x7ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 766247770432944429179173513575154591809369561091801088 (0x800000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 766247770432944429179173513575154591809369561091801089 (0x800000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1532495540865888858358347027150309183618739122183602175 (0xfffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1532495540865888858358347027150309183618739122183602176 (0x1000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1532495540865888858358347027150309183618739122183602177 (0x1000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3064991081731777716716694054300618367237478244367204351 (0x1fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3064991081731777716716694054300618367237478244367204352 (0x2000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3064991081731777716716694054300618367237478244367204353 (0x2000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 6129982163463555433433388108601236734474956488734408703 (0x3fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 6129982163463555433433388108601236734474956488734408704 (0x4000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 6129982163463555433433388108601236734474956488734408705 (0x4000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 12259964326927110866866776217202473468949912977468817407 (0x7fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 12259964326927110866866776217202473468949912977468817408 (0x8000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 12259964326927110866866776217202473468949912977468817409 (0x8000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 24519928653854221733733552434404946937899825954937634815 (0xffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 24519928653854221733733552434404946937899825954937634816 (0x10000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 24519928653854221733733552434404946937899825954937634817 (0x10000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 49039857307708443467467104868809893875799651909875269631 (0x1ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 49039857307708443467467104868809893875799651909875269632 (0x20000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 49039857307708443467467104868809893875799651909875269633 (0x20000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 98079714615416886934934209737619787751599303819750539263 (0x3ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 98079714615416886934934209737619787751599303819750539264 (0x40000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 98079714615416886934934209737619787751599303819750539265 (0x40000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 196159429230833773869868419475239575503198607639501078527 (0x7ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 196159429230833773869868419475239575503198607639501078528 (0x80000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 196159429230833773869868419475239575503198607639501078529 (0x80000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 392318858461667547739736838950479151006397215279002157055 (0xfffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 392318858461667547739736838950479151006397215279002157056 (0x100000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 392318858461667547739736838950479151006397215279002157057 (0x100000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 784637716923335095479473677900958302012794430558004314111 (0x1fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 784637716923335095479473677900958302012794430558004314112 (0x200000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 784637716923335095479473677900958302012794430558004314113 (0x200000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1569275433846670190958947355801916604025588861116008628223 (0x3fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1569275433846670190958947355801916604025588861116008628224 (0x400000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1569275433846670190958947355801916604025588861116008628225 (0x400000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3138550867693340381917894711603833208051177722232017256447 (0x7fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3138550867693340381917894711603833208051177722232017256448 (0x800000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3138550867693340381917894711603833208051177722232017256449 (0x800000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 6277101735386680763835789423207666416102355444464034512895 (0xffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 6277101735386680763835789423207666416102355444464034512896 (0x1000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 6277101735386680763835789423207666416102355444464034512897 (0x1000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 12554203470773361527671578846415332832204710888928069025791 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 12554203470773361527671578846415332832204710888928069025792 (0x2000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 12554203470773361527671578846415332832204710888928069025793 (0x2000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 25108406941546723055343157692830665664409421777856138051583 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 25108406941546723055343157692830665664409421777856138051584 (0x4000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 25108406941546723055343157692830665664409421777856138051585 (0x4000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 50216813883093446110686315385661331328818843555712276103167 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 50216813883093446110686315385661331328818843555712276103168 (0x8000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 50216813883093446110686315385661331328818843555712276103169 (0x8000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 100433627766186892221372630771322662657637687111424552206335 (0xfffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 100433627766186892221372630771322662657637687111424552206336 (0x10000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 100433627766186892221372630771322662657637687111424552206337 (0x10000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 200867255532373784442745261542645325315275374222849104412671 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 200867255532373784442745261542645325315275374222849104412672 (0x20000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 200867255532373784442745261542645325315275374222849104412673 (0x20000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 401734511064747568885490523085290650630550748445698208825343 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 401734511064747568885490523085290650630550748445698208825344 (0x40000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 401734511064747568885490523085290650630550748445698208825345 (0x40000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 803469022129495137770981046170581301261101496891396417650687 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 803469022129495137770981046170581301261101496891396417650688 (0x80000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 803469022129495137770981046170581301261101496891396417650689 (0x80000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1606938044258990275541962092341162602522202993782792835301375 (0xffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1606938044258990275541962092341162602522202993782792835301376 (0x100000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1606938044258990275541962092341162602522202993782792835301377 (0x100000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3213876088517980551083924184682325205044405987565585670602751 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3213876088517980551083924184682325205044405987565585670602752 (0x200000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3213876088517980551083924184682325205044405987565585670602753 (0x200000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 6427752177035961102167848369364650410088811975131171341205503 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 6427752177035961102167848369364650410088811975131171341205504 (0x400000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 6427752177035961102167848369364650410088811975131171341205505 (0x400000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 12855504354071922204335696738729300820177623950262342682411007 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 12855504354071922204335696738729300820177623950262342682411008 (0x800000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 12855504354071922204335696738729300820177623950262342682411009 (0x800000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 25711008708143844408671393477458601640355247900524685364822015 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 25711008708143844408671393477458601640355247900524685364822016 (0x1000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 25711008708143844408671393477458601640355247900524685364822017 (0x1000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 51422017416287688817342786954917203280710495801049370729644031 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 51422017416287688817342786954917203280710495801049370729644032 (0x2000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 51422017416287688817342786954917203280710495801049370729644033 (0x2000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 102844034832575377634685573909834406561420991602098741459288063 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 102844034832575377634685573909834406561420991602098741459288064 (0x4000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 102844034832575377634685573909834406561420991602098741459288065 (0x4000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 205688069665150755269371147819668813122841983204197482918576127 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 205688069665150755269371147819668813122841983204197482918576128 (0x8000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 205688069665150755269371147819668813122841983204197482918576129 (0x8000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 411376139330301510538742295639337626245683966408394965837152255 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 411376139330301510538742295639337626245683966408394965837152256 (0x10000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 411376139330301510538742295639337626245683966408394965837152257 (0x10000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 822752278660603021077484591278675252491367932816789931674304511 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 822752278660603021077484591278675252491367932816789931674304512 (0x20000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 822752278660603021077484591278675252491367932816789931674304513 (0x20000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1645504557321206042154969182557350504982735865633579863348609023 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1645504557321206042154969182557350504982735865633579863348609024 (0x40000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1645504557321206042154969182557350504982735865633579863348609025 (0x40000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3291009114642412084309938365114701009965471731267159726697218047 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3291009114642412084309938365114701009965471731267159726697218048 (0x80000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3291009114642412084309938365114701009965471731267159726697218049 (0x80000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 6582018229284824168619876730229402019930943462534319453394436095 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 6582018229284824168619876730229402019930943462534319453394436096 (0x100000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 6582018229284824168619876730229402019930943462534319453394436097 (0x100000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 13164036458569648337239753460458804039861886925068638906788872191 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 13164036458569648337239753460458804039861886925068638906788872192 (0x200000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 13164036458569648337239753460458804039861886925068638906788872193 (0x200000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 26328072917139296674479506920917608079723773850137277813577744383 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 26328072917139296674479506920917608079723773850137277813577744384 (0x400000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 26328072917139296674479506920917608079723773850137277813577744385 (0x400000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 52656145834278593348959013841835216159447547700274555627155488767 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 52656145834278593348959013841835216159447547700274555627155488768 (0x800000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 52656145834278593348959013841835216159447547700274555627155488769 (0x800000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 105312291668557186697918027683670432318895095400549111254310977535 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 105312291668557186697918027683670432318895095400549111254310977536 (0x1000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 105312291668557186697918027683670432318895095400549111254310977537 (0x1000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 210624583337114373395836055367340864637790190801098222508621955071 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 210624583337114373395836055367340864637790190801098222508621955072 (0x2000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 210624583337114373395836055367340864637790190801098222508621955073 (0x2000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 421249166674228746791672110734681729275580381602196445017243910143 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 421249166674228746791672110734681729275580381602196445017243910144 (0x4000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 421249166674228746791672110734681729275580381602196445017243910145 (0x4000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 842498333348457493583344221469363458551160763204392890034487820287 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 842498333348457493583344221469363458551160763204392890034487820288 (0x8000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 842498333348457493583344221469363458551160763204392890034487820289 (0x8000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1684996666696914987166688442938726917102321526408785780068975640575 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1684996666696914987166688442938726917102321526408785780068975640576 (0x10000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1684996666696914987166688442938726917102321526408785780068975640577 (0x10000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3369993333393829974333376885877453834204643052817571560137951281151 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3369993333393829974333376885877453834204643052817571560137951281152 (0x20000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3369993333393829974333376885877453834204643052817571560137951281153 (0x20000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 6739986666787659948666753771754907668409286105635143120275902562303 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 6739986666787659948666753771754907668409286105635143120275902562304 (0x40000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 6739986666787659948666753771754907668409286105635143120275902562305 (0x40000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 13479973333575319897333507543509815336818572211270286240551805124607 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 13479973333575319897333507543509815336818572211270286240551805124608 (0x80000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 13479973333575319897333507543509815336818572211270286240551805124609 (0x80000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 26959946667150639794667015087019630673637144422540572481103610249215 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 26959946667150639794667015087019630673637144422540572481103610249216 (0x100000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 26959946667150639794667015087019630673637144422540572481103610249217 (0x100000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 53919893334301279589334030174039261347274288845081144962207220498431 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 53919893334301279589334030174039261347274288845081144962207220498432 (0x200000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 53919893334301279589334030174039261347274288845081144962207220498433 (0x200000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 107839786668602559178668060348078522694548577690162289924414440996863 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 107839786668602559178668060348078522694548577690162289924414440996864 (0x400000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 107839786668602559178668060348078522694548577690162289924414440996865 (0x400000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 215679573337205118357336120696157045389097155380324579848828881993727 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 215679573337205118357336120696157045389097155380324579848828881993728 (0x800000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 215679573337205118357336120696157045389097155380324579848828881993729 (0x800000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 431359146674410236714672241392314090778194310760649159697657763987455 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 431359146674410236714672241392314090778194310760649159697657763987456 (0x1000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 431359146674410236714672241392314090778194310760649159697657763987457 (0x1000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 862718293348820473429344482784628181556388621521298319395315527974911 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 862718293348820473429344482784628181556388621521298319395315527974912 (0x2000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 862718293348820473429344482784628181556388621521298319395315527974913 (0x2000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1725436586697640946858688965569256363112777243042596638790631055949823 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1725436586697640946858688965569256363112777243042596638790631055949824 (0x4000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1725436586697640946858688965569256363112777243042596638790631055949825 (0x4000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3450873173395281893717377931138512726225554486085193277581262111899647 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3450873173395281893717377931138512726225554486085193277581262111899648 (0x8000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3450873173395281893717377931138512726225554486085193277581262111899649 (0x8000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 6901746346790563787434755862277025452451108972170386555162524223799295 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 6901746346790563787434755862277025452451108972170386555162524223799296 (0x10000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 6901746346790563787434755862277025452451108972170386555162524223799297 (0x10000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 13803492693581127574869511724554050904902217944340773110325048447598591 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 13803492693581127574869511724554050904902217944340773110325048447598592 (0x20000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 13803492693581127574869511724554050904902217944340773110325048447598593 (0x20000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 27606985387162255149739023449108101809804435888681546220650096895197183 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 27606985387162255149739023449108101809804435888681546220650096895197184 (0x40000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 27606985387162255149739023449108101809804435888681546220650096895197185 (0x40000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 55213970774324510299478046898216203619608871777363092441300193790394367 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 55213970774324510299478046898216203619608871777363092441300193790394368 (0x80000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 55213970774324510299478046898216203619608871777363092441300193790394369 (0x80000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 110427941548649020598956093796432407239217743554726184882600387580788735 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 110427941548649020598956093796432407239217743554726184882600387580788736 (0x100000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 110427941548649020598956093796432407239217743554726184882600387580788737 (0x100000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 220855883097298041197912187592864814478435487109452369765200775161577471 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 220855883097298041197912187592864814478435487109452369765200775161577472 (0x200000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 220855883097298041197912187592864814478435487109452369765200775161577473 (0x200000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 441711766194596082395824375185729628956870974218904739530401550323154943 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 441711766194596082395824375185729628956870974218904739530401550323154944 (0x400000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 441711766194596082395824375185729628956870974218904739530401550323154945 (0x400000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 883423532389192164791648750371459257913741948437809479060803100646309887 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 883423532389192164791648750371459257913741948437809479060803100646309888 (0x800000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 883423532389192164791648750371459257913741948437809479060803100646309889 (0x800000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1766847064778384329583297500742918515827483896875618958121606201292619775 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1766847064778384329583297500742918515827483896875618958121606201292619776 (0x1000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1766847064778384329583297500742918515827483896875618958121606201292619777 (0x1000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3533694129556768659166595001485837031654967793751237916243212402585239551 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3533694129556768659166595001485837031654967793751237916243212402585239552 (0x2000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3533694129556768659166595001485837031654967793751237916243212402585239553 (0x2000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 7067388259113537318333190002971674063309935587502475832486424805170479103 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 7067388259113537318333190002971674063309935587502475832486424805170479104 (0x4000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 7067388259113537318333190002971674063309935587502475832486424805170479105 (0x4000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 14134776518227074636666380005943348126619871175004951664972849610340958207 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 14134776518227074636666380005943348126619871175004951664972849610340958208 (0x8000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 14134776518227074636666380005943348126619871175004951664972849610340958209 (0x8000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 28269553036454149273332760011886696253239742350009903329945699220681916415 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 28269553036454149273332760011886696253239742350009903329945699220681916416 (0x10000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 28269553036454149273332760011886696253239742350009903329945699220681916417 (0x10000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 56539106072908298546665520023773392506479484700019806659891398441363832831 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 56539106072908298546665520023773392506479484700019806659891398441363832832 (0x20000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 56539106072908298546665520023773392506479484700019806659891398441363832833 (0x20000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 113078212145816597093331040047546785012958969400039613319782796882727665663 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 113078212145816597093331040047546785012958969400039613319782796882727665664 (0x40000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 113078212145816597093331040047546785012958969400039613319782796882727665665 (0x40000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 226156424291633194186662080095093570025917938800079226639565593765455331327 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 226156424291633194186662080095093570025917938800079226639565593765455331328 (0x80000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 226156424291633194186662080095093570025917938800079226639565593765455331329 (0x80000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 452312848583266388373324160190187140051835877600158453279131187530910662655 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 452312848583266388373324160190187140051835877600158453279131187530910662656 (0x100000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 452312848583266388373324160190187140051835877600158453279131187530910662657 (0x100000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 904625697166532776746648320380374280103671755200316906558262375061821325311 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 904625697166532776746648320380374280103671755200316906558262375061821325312 (0x200000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 904625697166532776746648320380374280103671755200316906558262375061821325313 (0x200000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650623 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650624 (0x400000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650625 (0x400000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301247 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301248 (0x800000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301249 (0x800000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602495 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602496 (0x1000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602497 (0x1000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b1111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204991 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204992 (0x2000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b10000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204993 (0x2000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b11111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409983 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409984 (0x4000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b100000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409985 (0x4000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0b111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111111'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819967 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819968 (0x8000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0b1000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000001'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819969 (0x8000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
diff --git a/src/Compilers/Z/Bounds/Interpretation.v b/src/Compilers/Z/Bounds/Interpretation.v
deleted file mode 100644
index e37b64a09..000000000
--- a/src/Compilers/Z/Bounds/Interpretation.v
+++ /dev/null
@@ -1,289 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Relations.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.ZRange.
-Require Import Crypto.Util.ZRange.Operations.
-Require Import Crypto.Util.ZUtil.Definitions.
-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.
- Definition interp_base_type (ty : base_type) : Set := t.
-
- Section ops.
- (** Generic helper definitions *)
- Definition t_map1 (f : Z -> Z) : t -> t
- := fun x => ZRange.two_corners f x.
- Definition t_map2 (f : Z -> Z -> Z) : t -> t -> t
- := fun x y => ZRange.four_corners f x y.
- Definition t_map3 (f : Z -> Z -> Z -> Z) : t -> t -> t -> t
- := fun x y z => ZRange.eight_corners f x y z.
- (** Definitions of the actual bounds propogation *)
- (** Rules for adding new operations:
-
- - Use [t_mapn] to if the underlying operation on [Z] is
- monotonic in all [n] of its arguments ([t_mapn] handles both
- monotonic non-increasing and monotonic non-decreasing) *)
-
- 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 max_abs_bound (x : t) : Z
- := upper (ZRange.abs x).
- Definition upper_lor_and_bounds (x y : Z) : Z
- := 2^(1 + Z.log2_up (Z.max x y)).
- Definition extreme_lor_land_bounds (x y : t) : t
- := let mx := max_abs_bound x in
- let my := max_abs_bound y in
- {| lower := -upper_lor_and_bounds mx my ; upper := upper_lor_and_bounds mx my |}.
- Definition extremization_bounds (f : t -> t -> t) (x y : t) : t
- := 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
- := extremization_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
- := extremization_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 zselect' (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 zselect (c r1 r2 : t) : t
- := zselect' r1 r2.
- Definition add_with_carry : t -> t -> t -> t
- := t_map3 Z.add_with_carry.
- Definition sub_with_borrow : t -> t -> t -> t
- := t_map3 Z.sub_with_borrow.
- Definition modulo_pow2_constant : Z -> t -> t
- := fun e x
- => let d := 2^e in
- let (l, u) := (lower x, upper x) in
- {| lower := if l / d =? u / d then Z.min (l mod d) (u mod d) else Z.min 0 (d + 1);
- upper := if l / d =? u / d then Z.max (l mod d) (u mod d) else Z.max 0 (d - 1) |}.
- Definition div_pow2_constant : Z -> t -> t
- := fun e x
- => let d := 2^e in
- let (l, u) := (lower x, upper x) in
- {| lower := l / d ; upper := u / d |}.
- Definition opp_div_pow2_constant : Z -> t -> t
- := fun e x
- => let d := 2^e in
- let (l, u) := (lower x, upper x) in
- {| lower := -(u / d) ; upper := -(l / d) |}.
- 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 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
- := 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
- := cmovle' r1 r2.
-
- Definition id_with_alt {T1 T2 Tout} (x : interp_base_type T1) (y : interp_base_type T2)
- : interp_base_type Tout
- := match T1, T2, Tout with
- | TZ, TZ, TZ => y
- | _, _, _ => x
- end.
- End ops.
- Section ops_with_carry.
- Context (carry_boundary_bit_width : Z).
- Definition get_carry : t -> t * t
- := fun v =>
- (modulo_pow2_constant carry_boundary_bit_width v,
- div_pow2_constant carry_boundary_bit_width v).
- Definition get_borrow : t -> t * t
- := fun v =>
- (modulo_pow2_constant carry_boundary_bit_width v,
- opp_div_pow2_constant carry_boundary_bit_width v).
- Definition add_with_get_carry : t -> t -> t -> t * t
- := fun c x y
- => get_carry (add_with_carry c x y).
- Definition sub_with_get_borrow : t -> t -> t -> t * t
- := fun c x y
- => get_borrow (sub_with_borrow c x y).
- Definition mul_split : t -> t -> t * t
- := fun x y => get_carry (mul x y).
- End ops_with_carry.
-
- 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 log_bit_width_of_base_type ty : option nat
- := match ty with
- | TZ => None
- | TWord logsz => Some logsz
- end.
-
- Definition bit_width_of_base_type ty : option Z
- := option_map (fun logsz => 2^Z.of_nat logsz)%Z (log_bit_width_of_base_type ty).
-
- Definition truncation_bounds' bit_width (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 truncation_bounds ty : interp_base_type ty -> interp_base_type ty
- := truncation_bounds' (bit_width_of_base_type ty).
-
- Definition interp_op {src dst} (f : op src dst)
- (x : interp_flat_type interp_base_type src)
- : interp_flat_type interp_base_type dst
- := SmartVarfMap
- truncation_bounds
- (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 _ => {| lower := v ; upper := v |}
- | Add _ _ T => fun xy => add (fst xy) (snd xy)
- | Sub _ _ T => fun xy => sub (fst xy) (snd xy)
- | Mul _ _ T => fun xy => mul (fst xy) (snd xy)
- | Shl _ _ T => fun xy => shl (fst xy) (snd xy)
- | Shr _ _ T => fun xy => shr (fst xy) (snd xy)
- | Land _ _ T => fun xy => land (fst xy) (snd xy)
- | Lor _ _ T => fun xy => lor (fst xy) (snd xy)
- | Opp _ T => fun x => opp x
- | IdWithAlt _ _ T => fun xy => id_with_alt (fst xy) (snd xy)
- | Zselect _ _ _ T => fun cxy => let '(c, x, y) := eta3 cxy in zselect c x y
- | MulSplit carry_boundary_bit_width _ _ T1 T2
- => fun xy => mul_split carry_boundary_bit_width (fst xy) (snd xy)
- | AddWithCarry _ _ _ T => fun cxy => let '(c, x, y) := eta3 cxy in add_with_carry c x y
- | AddWithGetCarry carry_boundary_bit_width _ _ _ T1 T2
- => fun cxy => let '(c, x, y) := eta3 cxy in
- add_with_get_carry carry_boundary_bit_width c x y
- | SubWithBorrow _ _ _ T => fun cxy => let '(c, x, y) := eta3 cxy in sub_with_borrow c x y
- | SubWithGetBorrow carry_boundary_bit_width _ _ _ T1 T2
- => fun cxy => let '(c, x, y) := eta3 cxy in
- sub_with_get_borrow carry_boundary_bit_width c x y
- end%bounds x).
-
- 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 smallest_logsz
- (round_up : nat -> option nat)
- (b : t)
- : option nat
- := if (0 <=? lower b)%Z
- then Some (Z.to_nat (Z.log2_up (Z.log2_up (1 + upper b))))
- else None.
- Definition actual_logsz
- (round_up : nat -> option nat)
- (b : t)
- : option nat
- := if (0 <=? lower b)%Z
- then let smallest_lgsz := (Z.to_nat (Z.log2_up (Z.log2_up (1 + upper b)))) in
- let lgsz := round_up smallest_lgsz in
- match lgsz with
- | Some lgsz
- => if Nat.leb smallest_lgsz lgsz
- then Some lgsz
- else None
- | None => None
- end
- else None.
- Definition bounds_to_base_type
- {round_up : nat -> option nat}
- (T : base_type)
- (b : interp_base_type T)
- : base_type
- := match T with
- | TZ => match actual_logsz round_up b with
- | Some lgsz => TWord lgsz
- | None => TZ
- end
- | TWord _
- => match smallest_logsz round_up b with
- | Some lgsz => TWord lgsz
- | None => TZ
- end
- end.
-
- Definition option_min (a : nat) (b : option nat)
- := match b with
- | Some b => Nat.min a b
- | None => a
- end.
-
- Definition round_up_to_in_list (allowable_lgsz : list nat)
- (lgsz : nat)
- : option nat
- := let good_lgsz := List.filter (Nat.leb lgsz) allowable_lgsz in
- List.fold_right (fun a b => Some (option_min a b)) None good_lgsz.
-
- Definition ComputeBounds {t} (e : Expr t)
- (input_bounds : interp_flat_type interp_base_type (domain t))
- : interp_flat_type interp_base_type (codomain t)
- := Compilers.Syntax.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 : forall {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/IsBoundedBy.v b/src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v
deleted file mode 100644
index b23e0ff1b..000000000
--- a/src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v
+++ /dev/null
@@ -1,214 +0,0 @@
-Require Import Coq.Classes.Morphisms.
-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.Z.Bounds.InterpretationLemmas.Tactics.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Util.ZRange.CornersMonotoneBounds.
-Require Import Crypto.Util.ZUtil.Stabilization.
-Require Import Crypto.Util.ZUtil.MulSplit.
-Require Import Crypto.Util.ZUtil.Modulo.
-Require Import Crypto.Util.ZUtil.LandLorShiftBounds.
-Require Import Crypto.Util.ZUtil.Morphisms.
-Require Import Crypto.Util.ZUtil.Definitions.
-Require Import Crypto.Util.ZUtil.Tactics.LtbToLt.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.FixedWordSizesEquality.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.UniquePose.
-Require Import Crypto.Util.Tactics.SpecializeBy.
-Require Import Crypto.Util.ZUtil.Tactics.SplitMinMax.
-
-Local Open Scope Z_scope.
-
-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 Bounds.log_bit_width_of_base_type option_map 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.
-
-Lemma is_bounded_by_compose T1 T2 f_v bs v f_bs fv
- (H : Bounds.is_bounded_by (T:=Tbase T1) bs v)
- (Hf : forall bs v, Bounds.is_bounded_by (T:=Tbase T1) bs v -> Bounds.is_bounded_by (T:=Tbase T2) (f_bs bs) (f_v v))
- (Hfv : f_v v = fv)
- : Bounds.is_bounded_by (T:=Tbase T2) (f_bs bs) fv.
-Proof.
- subst; eauto.
-Qed.
-
-Local Existing Instances Z.log2_up_le_Proper Z.add_le_Proper Z.sub_with_borrow_le_Proper.
-Lemma land_upper_lor_land_bounds a b
- : Z.abs (Z.land a b) <= Bounds.upper_lor_and_bounds (Z.abs a) (Z.abs b).
-Proof.
- unfold Bounds.upper_lor_and_bounds.
- apply stabilizes_bounded; auto with zarith.
- rewrite <- !Z.max_mono by exact _.
- apply land_stabilizes; apply stabilization_time_weaker.
-Qed.
-
-Lemma lor_upper_lor_land_bounds a b
- : Z.abs (Z.lor a b) <= Bounds.upper_lor_and_bounds (Z.abs a) (Z.abs b).
-Proof.
- unfold Bounds.upper_lor_and_bounds.
- apply stabilizes_bounded; auto with zarith.
- rewrite <- !Z.max_mono by exact _.
- apply lor_stabilizes; apply stabilization_time_weaker.
-Qed.
-
-Lemma upper_lor_and_bounds_Proper
- : Proper (Z.le ==> Z.le ==> Z.le) Bounds.upper_lor_and_bounds.
-Proof.
- intros ??? ???.
- unfold Bounds.upper_lor_and_bounds.
- auto with zarith.
-Qed.
-
-Local Arguments Z.pow !_ !_.
-Local Arguments Z.log2_up !_.
-Local Arguments Z.add !_ !_.
-Lemma land_bounds_extreme xb yb x y
- (Hx : ZRange.is_bounded_by' None xb x)
- (Hy : ZRange.is_bounded_by' None yb y)
- : ZRange.is_bounded_by' None (Bounds.extreme_lor_land_bounds xb yb) (Z.land x y).
-Proof.
- apply ZRange.monotonify2; auto;
- unfold Bounds.extreme_lor_land_bounds;
- [ apply land_upper_lor_land_bounds
- | apply upper_lor_and_bounds_Proper ].
-Qed.
-Lemma lor_bounds_extreme xb yb x y
- (Hx : ZRange.is_bounded_by' None xb x)
- (Hy : ZRange.is_bounded_by' None yb y)
- : ZRange.is_bounded_by' None (Bounds.extreme_lor_land_bounds xb yb) (Z.lor x y).
-Proof.
- apply ZRange.monotonify2; auto;
- unfold Bounds.extreme_lor_land_bounds;
- [ apply lor_upper_lor_land_bounds
- | apply upper_lor_and_bounds_Proper ].
-Qed.
-
-Local Arguments N.ldiff : simpl never.
-Local Arguments Z.pow : simpl never.
-Local Arguments Z.add !_ !_.
-Local Existing Instances Z.add_le_Proper Z.sub_le_flip_le_Proper Z.log2_up_le_Proper Z.pow_Zpos_le_Proper Z.sub_le_eq_Proper Z.add_with_carry_le_Proper.
-Local Hint Extern 1 => progress cbv beta iota : typeclass_instances.
-Local Ltac ibbio_prefin :=
- [ > | intros | simpl; reflexivity ].
-Local Ltac apply_is_bounded_by_truncation_bounds :=
- cbv [id
- Bounds.interp_op interp_op Bounds.is_bounded_by Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop SmartVarfMap smart_interp_flat_map lift_op SmartFlatTypeMapUnInterp cast_const Zinterp_op SmartFlatTypeMapInterp2
- Z.add_with_get_carry Bounds.add_with_get_carry Bounds.sub_with_get_borrow Bounds.get_carry Bounds.get_borrow Z.get_carry Bounds.mul_split];
- cbn [interpToZ fst snd];
- repeat match goal with
- | [ |- _ /\ _ ] => split
- end;
- lazymatch goal with
- | [ |- Bounds.is_bounded_by' (Bounds.truncation_bounds _ _) (ZToInterp _) ]
- => apply is_bounded_by_truncation_bounds'
- end.
-Local Ltac handle_mul :=
- apply (ZRange.monotone_four_corners_genb Z.mul); try (split; auto);
- unfold Basics.flip;
- let x := fresh "x" in
- intro x;
- exists (0 <=? x);
- break_match; Z.ltb_to_lt;
- intros ???; nia.
-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)
- (H_side : to_prop (interped_op_side_conditions opc v))
- : Bounds.is_bounded_by (Bounds.interp_op opc bs) (Syntax.interp_op _ _ opc v).
-Proof.
- destruct opc; apply_is_bounded_by_truncation_bounds;
- [ ..
- | eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v mod _)) (v:=ZToInterp _);
- [ .. | cbn -[Z.mul_split_at_bitwidth]; rewrite Z.mul_split_at_bitwidth_mod ];
- ibbio_prefin
- | eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v / _)) (v:=ZToInterp _);
- [ .. | cbn -[Z.mul_split_at_bitwidth]; rewrite Z.mul_split_at_bitwidth_div ];
- ibbio_prefin
- |
- | eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v mod _)) (v:=ZToInterp _);
- ibbio_prefin
- | eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v / _)) (v:=ZToInterp _);
- ibbio_prefin
- |
- | eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v mod _)) (v:=ZToInterp _);
- ibbio_prefin
- | eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (-(v / _))) (v:=ZToInterp _);
- ibbio_prefin ];
- repeat first [ progress simpl in *
- | progress cbv [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 ].
- { apply (@ZRange.monotone_four_corners true true Z.add _); split; auto. }
- { apply (@ZRange.monotone_four_corners true false Z.sub _); split; auto. }
- { handle_mul. }
- { apply (ZRange.monotone_four_corners_genb Z.shiftl); try (split; auto);
- [ eexists; apply Z.shiftl_le_Proper1
- | exists true; apply Z.shiftl_le_Proper2 ]. }
- { apply (ZRange.monotone_four_corners_genb Z.shiftr); try (split; auto);
- [ eexists; apply Z.shiftr_le_Proper1
- | exists true; apply Z.shiftr_le_Proper2 ]. }
- { cbv [Bounds.land Bounds.extremization_bounds]; break_innermost_match;
- [ apply land_bounds_extreme; split; auto | .. ];
- repeat first [ progress simpl in *
- | Zarith_t_step
- | rewriter_t
- | progress saturate_land_lor_facts
- | split_min_max; omega ]. }
- { cbv [Bounds.lor Bounds.extremization_bounds]; break_innermost_match;
- [ apply lor_bounds_extreme; split; auto | .. ];
- repeat first [ progress simpl in *
- | Zarith_t_step
- | rewriter_t
- | progress Zarith_land_lor_t_step
- | solve [ split_min_max; try omega; try Zarith_land_lor_t_step ] ]. }
- { repeat first [ progress destruct_head Bounds.t
- | progress simpl in *
- | progress split_min_max
- | omega ]. }
- { cbv [Bounds.id_with_alt];
- break_innermost_match; simpl in *; Z.ltb_to_lt; subst;
- split; assumption. }
- { destruct_head Bounds.t; cbv [Bounds.zselect' Z.zselect].
- break_innermost_match; split_min_max; omega. }
- { handle_mul. }
- { apply Z.mod_bound_min_max; auto. }
- { handle_mul. }
- { auto with zarith. }
- { apply (@ZRange.monotone_eight_corners true true true Z.add_with_carry _ _); split; auto. }
- { apply (@ZRange.monotone_eight_corners true true true Z.add_with_carry _ _); split; auto. }
- { apply Z.mod_bound_min_max; auto. }
- { apply (@ZRange.monotone_eight_corners true true true Z.add_with_carry _ _); split; auto. }
- { auto with zarith. }
- { apply (@ZRange.monotone_eight_corners false true false Z.sub_with_borrow _ _); split; auto. }
- { apply (@ZRange.monotone_eight_corners false true false Z.sub_with_borrow _ _); split; auto. }
- { apply Z.mod_bound_min_max; auto. }
- { apply (@ZRange.monotone_eight_corners false true false Z.sub_with_borrow _ _); split; auto. }
- { auto with zarith. }
-Qed.
diff --git a/src/Compilers/Z/Bounds/InterpretationLemmas/PullCast.v b/src/Compilers/Z/Bounds/InterpretationLemmas/PullCast.v
deleted file mode 100644
index 7ffe0beb1..000000000
--- a/src/Compilers/Z/Bounds/InterpretationLemmas/PullCast.v
+++ /dev/null
@@ -1,192 +0,0 @@
-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.Z.Syntax.Equality.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Bounds.Interpretation.
-Require Import Crypto.Compilers.Z.Bounds.InterpretationLemmas.Tactics.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Util.ZUtil.Morphisms.
-Require Import Crypto.Util.ZUtil.Log2.
-Require Import Crypto.Util.ZUtil.Pow2.
-Require Import Crypto.Util.ZUtil.Tactics.LtbToLt.
-Require Import Crypto.Util.ZUtil.Tactics.RewriteModSmall.
-Require Import Crypto.Util.ZUtil.Tactics.PeelLe.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.FixedWordSizesEquality.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.PointedProp.
-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 Open Scope Z_scope.
-
-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 Instances Z.pow_Zpos_le_Proper Z.log2_up_le_Proper.
-
-Section with_round_up.
- Context (is_bounded_by_interp_op
- : forall 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)
- (Hside : to_prop (interped_op_side_conditions opc v)),
- Bounds.is_bounded_by (Bounds.interp_op opc bs) (Syntax.interp_op _ _ opc v)).
- Context {round_up : nat -> option nat}.
-
- Local Notation pick_typeb := (@Bounds.bounds_to_base_type round_up) (only parsing).
- Local Notation pick_type v := (SmartFlatTypeMap pick_typeb v).
-
- Local Ltac t_small :=
- repeat first [ progress cbv [Bounds.bounds_to_base_type Bounds.smallest_logsz Bounds.actual_logsz] in *
- | progress simpl in *
- | progress autorewrite with push_Zof_nat zsimplify_const
- | solve [ trivial ]
- | progress change (@interpToZ) with (fun t1 => cast_const (t1:=t1) (t2:=TZ)) in *
- | progress change (@cast_const TZ TZ) with (fun x : Z => x) in *
- | progress Z.ltb_to_lt
- | rewrite Z2Nat.id in * by auto with zarith
- | rewrite <- Z.log2_up_le_pow2_full in * by auto with zarith
- | break_innermost_match_step
- | apply conj
- | omega
- | rewrite <- Z.log2_up_le_full
- | match goal with
- | [ |- ?x < ?y ] => cut (x + 1 <= y); [ omega | ]
- | [ H : (_ <=? _)%nat = true |- _ ] => apply Nat.leb_le in H
- | [ H : (_ <= _)%nat |- _ ] => apply inj_le in H
- end ].
-
- Local Arguments cast_const : simpl never.
- Lemma interpToZ_cast_const_small T0 (bs : Bounds.interp_base_type T0)
- (v : interp_base_type (@Bounds.bounds_to_base_type round_up T0 bs))
- (H : match Bounds.bit_width_of_base_type T0 with
- | Some sz => 0 <= ZRange.lower bs /\ ZRange.upper bs < 2 ^ sz
- | None => True
- end)
- (Hv : ZRange.lower bs <= interpToZ (cast_const (t2:=T0) v) <= ZRange.upper bs)
- : interpToZ (cast_const (t2:=T0) v) = interpToZ v.
- Proof.
- destruct_head base_type; simpl in *; try reflexivity.
- unfold Bounds.smallest_logsz, Bounds.interp_base_type, cast_const in *.
- break_innermost_match_hyps; Z.ltb_to_lt; simpl in *;
- [ pose proof (wordToZ_range v) | omega ].
- rewrite wordToZ_ZToWord_mod_full.
- rewrite wordToZ_ZToWord_mod_full in Hv.
- revert Hv; apply Z.max_case_strong; Z.rewrite_mod_small; intros; try omega; [].
- rewrite Z.mod_small; [ reflexivity | split; auto with zarith ].
- autorewrite with push_Zof_nat zsimplify_const in *.
- rewrite Z2Nat.id in * by auto with zarith.
- destruct_head' and.
- eapply Z.lt_le_trans; [ eassumption | ].
- repeat first [ progress Z.peel_le
- | rewrite <- Z.log2_up_pow2 by auto with zarith; progress Z.peel_le
- | omega ].
- Qed.
-
- Local Existing Instances Z.pow_Zpos_le_Proper Z.pow_Zpos_le_Proper_flip Z.lt_le_flip_Proper_flip_impl.
- Lemma ZToInterp_cast_const_small T (bs : Bounds.interp_base_type T)
- v
- (H : match Bounds.bit_width_of_base_type T with
- | Some sz => 0 <= ZRange.lower bs /\ ZRange.upper bs < 2 ^ sz
- | None => True
- end)
- (Hb : ZRange.lower bs <= interpToZ (@ZToInterp T v) <= ZRange.upper bs)
- : @cast_const (@Bounds.bounds_to_base_type round_up T bs) T (ZToInterp v) = ZToInterp v.
- Proof.
- apply ZToInterp_eq_inj.
- rewrite ?interpToZ_ZToInterp_mod.
- rewrite interpToZ_ZToInterp_mod in Hb.
- destruct T; simpl in *.
- { unfold Bounds.actual_logsz.
- break_innermost_match; Z.ltb_to_lt; try reflexivity; [].
- t_small.
- apply Z.max_case_strong; intros; Z.rewrite_mod_small; omega. }
- { unfold Bounds.smallest_logsz.
- break_innermost_match_step; Z.ltb_to_lt; try omega; [].
- revert Hb; apply (Z.max_case_strong 0 v); intros; Z.rewrite_mod_small; try reflexivity.
- rewrite Z.max_r by auto with zarith.
- autorewrite with push_Zof_nat zsimplify_const in *.
- rewrite Z2Nat.id by auto with zarith.
- rewrite Z.mod_mod_small; try apply conj; auto with zarith;
- repeat first [ rewrite <- Z.log2_up_le_pow2_full in * by auto with zarith
- | rewrite <- Z.log2_up_le_full
- | omega
- | apply conj
- | progress Z.peel_le
- | rewrite <- Z.log2_up_pow2 by auto with zarith; progress Z.peel_le
- | match goal with
- | [ |- 2^?x mod 2^?y = 0 ]
- => destruct (Z.pow2_lt_or_divides x y);
- [ solve [ auto with zarith ]
- | exfalso; assert (2^y <= 2^x)
- | assumption ]
- end ]. }
- Qed.
-
- 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))
- (Hside : to_prop (interped_op_side_conditions opc (cast_back_flat_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); simpl; intros.
- repeat (destruct_one_head flat_type; try solve [ inversion opc ]);
- repeat first [ progress simpl in *
- | progress destruct_head unit
- | progress destruct_head True
- | progress cbv [unzify_op cast_back_flat_const lift_op Bounds.is_bounded_by' ZRange.is_bounded_by'] in *
- | progress change (@interpToZ TZ) with (fun x : Z => x) in *
- | progress specialize_by auto
- | progress specialize_by constructor
- | match goal with
- | [ H : forall v, _ <= ?f v <= _ /\ _ -> _ |- context[?f ?v'] ]
- => specialize (H v')
- | _ => progress specialize_by_assumption
- | _ => progress specialize_by auto
- | [ H : forall v : unit, _ |- _ ] => specialize (H tt)
- | [ H : forall v : _ * _, _ /\ _ -> _ |- _ ]
- => specialize (fun v1 v2 p1 p2 => H (v1, v2) (conj p1 p2));
- cbn [fst snd] in H;
- specialize (fun v1 p1 v2 p2 => H v1 v2 p1 p2)
- end
- | progress destruct_head'_and ];
- [
- | rewrite cast_const_idempotent_small by t_small; reflexivity
- | .. ];
- repeat match goal with
- | _ => progress change (@cast_const TZ) with @ZToInterp in *
- | [ |- context[@cast_const ?x TZ] ]
- => progress change (@cast_const x TZ) with (@interpToZ x) in *
- | [ H : context[@cast_const ?x TZ] |- _ ]
- => progress change (@cast_const x TZ) with (@interpToZ x) in *
- end;
- repeat match goal with
- | _ => reflexivity
- | _ => progress rewrite ?interpToZ_cast_const_small, ?ZToInterp_cast_const_small by auto
- | [ H : context[ZToInterp (?f (interpToZ (cast_const ?v)))] |- _ ]
- => rewrite (@interpToZ_cast_const_small _ _ v) in H by auto
- | [ H : context[(interpToZ (cast_const ?v), _)] |- _ ]
- => rewrite (@interpToZ_cast_const_small _ _ v) in H by auto
- | [ H : context[(_, interpToZ (cast_const ?v))] |- _ ]
- => rewrite (@interpToZ_cast_const_small _ _ v) in H by auto
- end.
- Qed.
-End with_round_up.
diff --git a/src/Compilers/Z/Bounds/InterpretationLemmas/Tactics.v b/src/Compilers/Z/Bounds/InterpretationLemmas/Tactics.v
deleted file mode 100644
index 6486b2e00..000000000
--- a/src/Compilers/Z/Bounds/InterpretationLemmas/Tactics.v
+++ /dev/null
@@ -1,197 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.micromega.Psatz.
-Require Import Crypto.Compilers.Z.Bounds.Interpretation.
-Require Import Crypto.Util.ZUtil.Hints.
-Require Import Crypto.Util.ZUtil.Hints.Core.
-Require Import Crypto.Util.ZUtil.ZSimplify.Core.
-Require Import Crypto.Util.ZUtil.Log2.
-Require Import Crypto.Util.ZUtil.Shift.
-Require Import Crypto.Util.ZUtil.LandLorShiftBounds.
-Require Import Crypto.Util.ZUtil.Tactics.LtbToLt.
-Require Import Crypto.Util.ZUtil.Notations.
-Require Import Crypto.Util.ZRange.Operations.
-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.
-Require Import Crypto.Util.ZUtil.Tactics.SplitMinMax.
-
-Local Open Scope Z_scope.
-
-Ltac break_t_step :=
- first [ progress destruct_head'_and
- | progress destruct_head'_or
- | progress destruct_head'_prod
- | progress split_andb
- | break_innermost_match_step ].
-
-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 ].
-
-Ltac specializer_t_step :=
- first [ progress specialize_by_assumption
- | progress specialize_by fin_t ].
-
-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 ].
-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.
-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.
-
-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.
-Ltac case_Zvar_nonneg_step :=
- match goal with
- | [ |- context[?x] ]
- => case_Zvar_nonneg_on x
- end.
-Ltac case_Zvar_nonneg := repeat case_Zvar_nonneg_step.
-
-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.
-
-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 >> ?y'] ]
- => unique assert (x >> y' <= y >> x') by (Z.shiftr_le_mono; omega)
- | [ H : ?x <= ?y, H' : ?x' <= ?y' |- context[?y >> ?x'] ]
- => unique assert (x >> y' <= y >> x') by (apply Z.shiftr_le_mono; omega)
- end.
-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 >> y' <= y >> x') by (apply Z.shiftr_le_mono; omega)
- end.
-Ltac preprocess_shift_min_max :=
- repeat first [ rewrite (Z.min_r (_ >> _) (_ >> _)) by (apply Z.shiftr_le_mono; omega)
- | rewrite (Z.min_l (_ >> _) (_ >> _)) by (apply Z.shiftr_le_mono; omega)
- | rewrite (Z.max_r (_ >> _) (_ >> _)) by (apply Z.shiftr_le_mono; omega)
- | rewrite (Z.max_l (_ >> _) (_ >> _)) by (apply Z.shiftr_le_mono; omega)
- | rewrite (Z.min_r (_ << _) (_ << _)) by (apply Z.shiftl_le_mono; omega)
- | rewrite (Z.min_l (_ << _) (_ << _)) by (apply Z.shiftl_le_mono; omega)
- | rewrite (Z.max_r (_ << _) (_ << _)) by (apply Z.shiftl_le_mono; omega)
- | rewrite (Z.max_l (_ << _) (_ << _)) by (apply Z.shiftl_le_mono; omega) ].
-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.
-Ltac handle_shift_neg :=
- repeat first [ rewrite !Z.shiftr_opp_r
- | rewrite !Z.shiftl_opp_r
- | rewrite !Z.shiftr_opp_l
- | rewrite !Z.shiftl_opp_l ].
-
-Ltac handle_four_corners_step_fast :=
- first [ progress destruct_head Bounds.t
- | progress cbv [ZRange.four_corners] in *
- | progress subst
- | Zarith_t_step
- | progress split_min_max
- | omega
- | nia ].
-Ltac handle_four_corners_step :=
- first [ handle_four_corners_step_fast
- | remove_binary_operation_le_hyps_step ].
-Ltac handle_four_corners :=
- lazymatch goal with
- | [ |- (ZRange.lower (ZRange.four_corners _ _ _) <= _ <= _)%Z ]
- => idtac
- end;
- repeat handle_four_corners_step.
-
-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 ].
diff --git a/src/Compilers/Z/Bounds/MapCastByDeBruijn.v b/src/Compilers/Z/Bounds/MapCastByDeBruijn.v
deleted file mode 100644
index c08ae1298..000000000
--- a/src/Compilers/Z/Bounds/MapCastByDeBruijn.v
+++ /dev/null
@@ -1,24 +0,0 @@
-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 (round_up : nat -> option nat)
- {t : type base_type}.
-
- Definition MapCastCompile := @MapCastCompile t.
- Definition MapCastDoCast
- := @MapCastDoCast
- (@Bounds.interp_base_type) (@Bounds.interp_op)
- (@Bounds.bounds_to_base_type round_up)
- (fun _ _ opc _ => @genericize_op _ _ _ opc _ _ _)
- t.
- Definition MapCastDoInterp
- := @MapCastDoInterp
- (@Bounds.interp_base_type) (@Bounds.bounds_to_base_type round_up)
- 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
deleted file mode 100644
index 1217cd721..000000000
--- a/src/Compilers/Z/Bounds/MapCastByDeBruijnInterp.v
+++ /dev/null
@@ -1,48 +0,0 @@
-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.InterpSideConditions.
-Require Import Crypto.Compilers.Z.Bounds.Interpretation.
-Require Import Crypto.Compilers.Z.Bounds.InterpretationLemmas.IsBoundedBy.
-Require Import Crypto.Compilers.Z.Bounds.InterpretationLemmas.PullCast.
-Require Import Crypto.Compilers.Z.Bounds.MapCastByDeBruijn.
-Require Import Crypto.Util.PointedProp.
-
-Lemma MapCastCorrect
- {round_up}
- {t} (e : Expr t)
- (Hwf : Wf e)
- (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
- : forall {b} e' (He':MapCast round_up e input_bounds = Some (existT _ b e'))
- v v' (Hv : Bounds.is_bounded_by input_bounds v /\ cast_back_flat_const v' = v)
- (Hside : to_prop (InterpSideConditions e v)),
- Compilers.Syntax.Interp (@Bounds.interp_op) e input_bounds = b
- /\ Bounds.is_bounded_by b (Compilers.Syntax.Interp interp_op e v)
- /\ cast_back_flat_const (Compilers.Syntax.Interp interp_op e' v') = (Compilers.Syntax.Interp interp_op e v).
-Proof.
- apply MapCastCorrect; auto.
- { apply is_bounded_by_interp_op. }
- { apply pull_cast_genericize_op, is_bounded_by_interp_op. }
-Qed.
-
-Lemma MapCastCorrect_eq
- {round_up}
- {t} (e : Expr t)
- {evb ev}
- (Hwf : Wf e)
- (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
- : forall {b} e' (He':MapCast round_up e input_bounds = Some (existT _ b e'))
- v v' (Hv : Bounds.is_bounded_by input_bounds v /\ cast_back_flat_const v' = v)
- (Hside : to_prop (InterpSideConditions e v))
- (Hevb : evb = Compilers.Syntax.Interp (@Bounds.interp_op) e input_bounds)
- (Hev : ev = Compilers.Syntax.Interp interp_op e v),
- evb = b
- /\ Bounds.is_bounded_by b ev
- /\ cast_back_flat_const (Compilers.Syntax.Interp interp_op e' v') = ev.
-Proof.
- intros; subst; apply MapCastCorrect; auto.
-Qed.
diff --git a/src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v b/src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v
deleted file mode 100644
index fabb8b2b8..000000000
--- a/src/Compilers/Z/Bounds/MapCastByDeBruijnWf.v
+++ /dev/null
@@ -1,42 +0,0 @@
-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.MapCastByDeBruijn.
-
-Definition Wf_MapCast
- {round_up}
- {t} (e : Expr t)
- (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
- {b} e' (He' : MapCast round_up e input_bounds = Some (existT _ b e'))
- (Hwf : Wf e)
- : Wf e'
- := @Wf_MapCast
- (@Bounds.interp_base_type) (@Bounds.interp_op)
- (@Bounds.bounds_to_base_type round_up)
- (fun _ _ opc _ => @genericize_op _ _ _ opc _ _ _)
- t e input_bounds b e' He' Hwf.
-
-Definition Wf_MapCast_arrow
- {round_up}
- {s d} (e : Expr (Arrow s d))
- (input_bounds : interp_flat_type Bounds.interp_base_type s)
- {b} e' (He' : MapCast round_up e input_bounds = Some (existT _ b e'))
- (Hwf : Wf e)
- : Wf e'
- := @Wf_MapCast_arrow
- (@Bounds.interp_base_type) (@Bounds.interp_op)
- (@Bounds.bounds_to_base_type round_up)
- (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
deleted file mode 100644
index c7060ea44..000000000
--- a/src/Compilers/Z/Bounds/Pipeline.v
+++ /dev/null
@@ -1,56 +0,0 @@
-(** * Reflective Pipeline *)
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Z.Bounds.Pipeline.Glue.
-Require Import Crypto.Compilers.Z.Bounds.Pipeline.ReflectiveTactics.
-Import ListNotations.
-Local Open Scope nat_scope.
-Local Open Scope list_scope.
-(** 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.
- Existing Instance DefaultedTypes.by_default.
-End Exports.
-
-Ltac refine_reflectively_gen allowable_bit_widths opts :=
- refine_to_reflective_glue allowable_bit_widths;
- do_reflective_pipeline opts.
-
-Ltac refine_reflectively128_with opts := refine_reflectively_gen [128; 256] opts.
-Ltac refine_reflectively64_with opts := refine_reflectively_gen [64; 128] opts.
-Ltac refine_reflectively32_with opts := refine_reflectively_gen [32; 64] opts.
-Ltac refine_reflectively128_with_bool_with opts := refine_reflectively_gen [1; 128; 256] opts.
-Ltac refine_reflectively64_with_bool_with opts := refine_reflectively_gen [1; 64; 128] opts.
-Ltac refine_reflectively32_with_bool_with opts := refine_reflectively_gen [1; 32; 64] opts.
-Ltac refine_reflectively128_with_uint8_with opts := refine_reflectively_gen [8; 128; 256] opts.
-Ltac refine_reflectively64_with_uint8_with opts := refine_reflectively_gen [8; 64; 128] opts.
-Ltac refine_reflectively32_with_uint8_with opts := refine_reflectively_gen [8; 32; 64] opts.
-Ltac refine_reflectively128 := refine_reflectively128_with default_PipelineOptions.
-Ltac refine_reflectively64 := refine_reflectively64_with default_PipelineOptions.
-Ltac refine_reflectively32 := refine_reflectively32_with default_PipelineOptions.
-Ltac refine_reflectively128_with_bool := refine_reflectively128_with_bool_with default_PipelineOptions.
-Ltac refine_reflectively64_with_bool := refine_reflectively64_with_bool_with default_PipelineOptions.
-Ltac refine_reflectively32_with_bool := refine_reflectively32_with_bool_with default_PipelineOptions.
-Ltac refine_reflectively128_with_uint8 := refine_reflectively128_with_uint8_with default_PipelineOptions.
-Ltac refine_reflectively64_with_uint8 := refine_reflectively64_with_uint8_with default_PipelineOptions.
-Ltac refine_reflectively32_with_uint8 := refine_reflectively32_with_uint8_with default_PipelineOptions.
-
-(** Convenience notations for options *)
-Definition anf := {| Pipeline.Definition.anf := true |}.
-Definition anf_without_adc_fusion := {| Pipeline.Definition.anf := true ; Pipeline.Definition.adc_fusion := false |}.
-Definition adc_fusion := {| Pipeline.Definition.adc_fusion := true |}.
-Definition default := default_PipelineOptions.
-
-(** The default *)
-Ltac refine_reflectively_with_bool_with opts := refine_reflectively64_with_bool_with opts.
-Ltac refine_reflectively_with_uint8_with opts := refine_reflectively64_with_uint8_with opts.
-Ltac refine_reflectively_with opts := refine_reflectively64_with opts.
-Ltac refine_reflectively := refine_reflectively_with default.
diff --git a/src/Compilers/Z/Bounds/Pipeline/Definition.v b/src/Compilers/Z/Bounds/Pipeline/Definition.v
deleted file mode 100644
index 435ec793e..000000000
--- a/src/Compilers/Z/Bounds/Pipeline/Definition.v
+++ /dev/null
@@ -1,285 +0,0 @@
-(** * 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.Linearize.
-Require Import Crypto.Compilers.LinearizeInterp.
-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 t) : Expr _
- := ExprEta (SimplifyArith false (Linearize 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 t)
- : forall x, Interp (PreWfPipeline e) x = Interp 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.CommonSubexpressionElimination.
-Require Import Crypto.Compilers.Z.CommonSubexpressionEliminationInterp.
-Require Import Crypto.Compilers.Z.CommonSubexpressionEliminationWf.*)
-Require Import Crypto.Compilers.Z.ArithmeticSimplifierWf.
-Require Import Crypto.Compilers.Z.RewriteAddToAdc.
-Require Import Crypto.Compilers.Z.RewriteAddToAdcWf.
-Require Import Crypto.Compilers.Z.RewriteAddToAdcInterp.
-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.
-Require Import Crypto.Util.DefaultedTypes.
-
-(** *** Definition of the Post-Wf Pipeline *)
-(** We define the record that holds various options to customize the
- pipeline. *)
-Record PipelineOptions :=
- {
- anf : with_default bool false;
- adc_fusion : with_default bool true;
- }.
-Definition default_PipelineOptions := {| anf := _ |}.
-
-(** Do not change the name or the type of these two definitions *)
-(** The definition [PostWfPreBoundsPipeline] is for the part of the
- pipeline that comes before [MapCast]; it must preserve the type of
- the expression. *)
-Definition PostWfPreBoundsPipeline
- (opts : PipelineOptions)
- {t} (e : Expr t)
- : Expr t
- := let e := InlineConst e in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := InlineConst (Linearize (SimplifyArith false e)) in
- let e := if opts.(anf) then InlineConst (ANormal e) else e in
- let e := if opts.(adc_fusion) then RewriteAdc e else e in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := if opts.(anf) then InlineConstAndOpp (ANormal e) else e in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- let e := InlineConstAndOpp (Linearize (SimplifyArith true e)) in
- (*let e := CSE false e in*)
- e.
-(** The definition [PostWfBoundsPipeline] is for the part of the
- pipeline that begins with [MapCast]. *)
-Definition PostWfBoundsPipeline
- round_up
- (opts : PipelineOptions)
- {t} (e0 : Expr t)
- (e : Expr t)
- (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
- : option (@ProcessedReflectivePackage round_up)
- := Build_ProcessedReflectivePackage_from_option_sigma
- e0 input_bounds
- (let e := MapCast _ e input_bounds in
- option_map
- (projT2_map
- (fun b e'
- => let e' := InlineConst e' in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := InlineConst (Linearize (SimplifyArith false e')) in
- let e' := if opts.(anf) then InlineConst (ANormal e') else e' in
- let e' := ExprEta e' in
- e'))
- e).
-
-(** *** Correctness proof of the Post-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.Compilers.Z.InterpSideConditions.
-Require Import Crypto.Compilers.InterpRewriting.
-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.
-Require Import Crypto.Util.PointedProp.
-
-Section with_round_up_list.
- Context {allowable_lgsz : list nat}.
-
- Local Notation pick_typeb := (@Bounds.bounds_to_base_type (Bounds.round_up_to_in_list allowable_lgsz)) (only parsing).
- Local Notation pick_type v := (SmartFlatTypeMap pick_typeb v).
- Local Opaque to_prop InterpSideConditions.
-
- Definition PostWfPreBoundsPipelineCorrect
- opts
- {t}
- (e : Expr t)
- (Hwf : Wf e)
- (e' := PostWfPreBoundsPipeline opts e)
- : (forall v, Interp e' v = Interp e v) /\ Wf e'.
- Proof using Type.
- (** These first two lines probably shouldn't change much *)
- unfold PostWfPreBoundsPipeline in *; subst e'.
- break_match_hyps.
- (** Now handle all the transformations that come before the word-size selection *)
- rewrite_reflective_interp_cached.
- split; intros; finish_rewrite_reflective_interp_cached; auto.
- Qed.
-
- Definition PostWfBoundsPipelineCorrect
- opts
- {t}
- (e0 : Expr t)
- (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
- (e : Expr t)
- (Hwf : Wf e)
- {b e'} (He' : PostWfBoundsPipeline _ opts e0 e input_bounds
- = Some {| input_expr := e0 ; 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)
- (Hside : to_prop (InterpSideConditions e v))
- : Bounds.is_bounded_by b (Interp e v)
- /\ cast_back_flat_const (Interp e' v') = Interp e v.
- Proof using Type.
- (** These first two lines probably shouldn't change much *)
- unfold PostWfBoundsPipeline, 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 *)
- all:rewrite_reflective_interp_cached.
- (** Now handle word-size selection *)
- all:(eapply MapCastCorrect_eq; [ | eassumption | eassumption | assumption | .. ];
- [ solve_wf_side_condition | reflexivity | ]).
- all:reflexivity.
- Qed.
-
- Definition PostWfPipelineCorrect
- opts
- {t}
- (e : Expr t)
- (input_bounds : interp_flat_type Bounds.interp_base_type (domain t))
- (Hwf : Wf e)
- (e' := PostWfPreBoundsPipeline opts e)
- {b e''} (He' : PostWfBoundsPipeline _ opts e 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)
- (Hside : to_prop (InterpSideConditions e' v))
- : Bounds.is_bounded_by b (Interp e v)
- /\ cast_back_flat_const (Interp e'' v') = Interp e v.
- Proof.
- rewrite <- (proj1 (PostWfPreBoundsPipelineCorrect opts e Hwf)) by assumption.
- eapply PostWfBoundsPipelineCorrect; eauto.
- apply PostWfPreBoundsPipelineCorrect; assumption.
- Qed.
-End with_round_up_list.
-
-(** ** 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 bbv.WordScope.
-
-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
deleted file mode 100644
index ed9abf9a6..000000000
--- a/src/Compilers/Z/Bounds/Pipeline/Glue.v
+++ /dev/null
@@ -1,550 +0,0 @@
-(** * 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.Logic.ImplAnd.
-Require Import Crypto.Util.Tactics.EvarExists.
-Require Import Crypto.Util.Tactics.GetGoal.
-Require Import Crypto.Util.Tactics.PrintContext.
-Require Import Crypto.Util.Tactics.MoveLetIn.
-Require Import Crypto.Util.Tactics.ClearAll.
-Require Import Crypto.Util.Tactics.ClearbodyAll.
-
-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 [goal_dlet_to_context_curried] moves to the
- context any [dlet x := y in ...] in the goal, curries each such
- moved definition, and then reifies the type of each such context
- variable. *)
-Ltac goal_dlet_to_context_curried :=
- lazymatch goal with
- | [ |- context[@Let_In ?A ?B ?x _] ]
- => let f := fresh in
- goal_dlet_to_context_step f;
- change_with_curried f;
- change_to_reified_type f;
- goal_dlet_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];
- goal_dlet_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;
- lazymatch goal with
- | [ |- context[@proj1_sig ?A ?P a] ]
- => pattern (@proj1_sig A P a)
- end;
- lazymatch goal with
- | [ |- ?P ?p1a ]
- => cut (dlet p := P in p p1a);
- [ repeat (clear_all; clearbody_all);
- 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 take a goal either 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
->>
-
- The tactic [maybe_reassoc_sig_and_eexists] also supports goals that
- don't need to be reassociated.
- *)
-Ltac reassoc_sig_and_eexists :=
- pattern_sig_sig_assoc;
- evar_exists.
-Ltac maybe_reassoc_sig_and_eexists _ :=
- lazymatch goal with
- | [ |- { a : ?T | _ } ]
- => lazymatch (eval hnf in T) with
- | { a' | _ }
- => reassoc_sig_and_eexists
- | _ => evar_exists
- end
- end.
-
-(** ** [intros_under_and] *)
-(** The [intros_under_and] tactic takes a goal of the form
-<<
-(A -> B -> ... -> Y -> Z) /\ (A -> B -> ... -> Y -> Z')
->>
- and turns it into a goal of the form
-<<
-Z /\ Z'
->>
- where [A], [B], ..., [Y] have been introduced into the context. *)
-Ltac intros_under_and _ :=
- repeat (apply (proj1 impl_and_iff); intro); intros.
-
-(** ** [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:(Tuple.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_map_wordToZ lvl descr top_subterm subterm :=
- lazymatch subterm with
- | wordToZ => idtac
- | Tuple.map ?f => check_is_map_wordToZ lvl descr top_subterm f
- | _ => let map_shape := uconstr:(Tuple.map) in
- let wordToZ_shape := uconstr:(wordToZ) in
- fail lvl descr
- "which are a repeated application of" map_shape "to" wordToZ_shape
- "but in the subterm" top_subterm "the function" subterm
- "was found, which is not of this form"
- end.
-Tactic Notation "check_is_map_wordToZ" int_or_var(lvl) string(descr) constr(term)
- := check_is_map_wordToZ lvl descr term term.
-
-Ltac check_is_bounded_by_shape subterm_type :=
- lazymatch subterm_type with
- | ZRange.is_bounded_by None ?bounds (Tuple.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 (Tuple.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
- | Tuple.map ?f ?fW
- => check_is_map_wordToZ 0 "The left-hand side of the second conjunct of the goal must be a tuple of terms" f;
- 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:(Tuple.map wordToZ ?fW) in
- let ef := uconstr:(?f) in
- let shape' := uconstr:(Tuple.map ?f ?fW) in
- let map_shape := uconstr:(Tuple.map) in
- let wordToZ_shape := uconstr:(wordToZ) 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
- "or" shape' "where" ef "is a repeated application of" map_shape "to" wordToZ_shape
- "but a subterm not matching this shape was found:" subterm
- end.
-Ltac check_RHS_Z_shape_rec subterm :=
- lazymatch subterm with
- | Tuple.map ?f ?fW
- => check_is_map_wordToZ
- 0
- "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" f
- | (?A, ?B)
- => check_RHS_Z_shape_rec A;
- check_RHS_Z_shape_rec B
- | _ => let G := get_goal in
- let shape := uconstr:(Tuple.map wordToZ ?fW) in
- let ef := uconstr:(?f) in
- let shape' := uconstr:(Tuple.map ?f ?fW) in
- let map_shape := uconstr:(Tuple.map) in
- let wordToZ_shape := uconstr:(wordToZ) 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
- "or" shape' "where" ef "is a repeated application of" map_shape "to" wordToZ_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:(Tuple.map wordToZ ?fW) in
- let ef := uconstr:(?f) in
- let shape' := uconstr:(Tuple.map ?f ?fW) in
- let map_shape := uconstr:(Tuple.map) in
- let wordToZ_shape := uconstr:(wordToZ) 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
- "or" shape' "where" ef "is a repeated application of" map_shape "to" wordToZ_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;
- let destruct_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 ] in
- let destruct_pair 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) in *;
- change (snd x) with (let (a, b) := x in b) in *;
- instantiate (1:=ltac:(destruct x)); destruct x ];
- (cbv beta iota) in
- let destruct_sig_or_pair v :=
- match v with
- | context[proj1_sig ?x] => destruct_sig x
- | context[fst ?x] => destruct_pair x
- | context[snd ?x] => destruct_pair x
- end in
- repeat match goal with
- | [ |- context[Tuple.map ?f ?v] ]
- => check_is_map_wordToZ 0 "DEBUG" f; destruct_sig_or_pair v
- | [ H := context[Tuple.map ?f ?v] |- _ ]
- => check_is_map_wordToZ 0 "DEBUG" f; destruct_sig_or_pair v
- | [ H : context[Tuple.map ?f ?v] |- _ ]
- => check_is_map_wordToZ 0 "DEBUG" f; destruct_sig_or_pair v
- | [ H : _ /\ _ |- _ ]
- => destruct H
- 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'))
- | Tuple.map ?f ?x
- => let dummy := match goal with
- | _ => check_is_map_wordToZ 1 "In unmap_wordToZ_tuple, expected terms" f
- end in
- 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 ?f ?arg
- => first [ check_is_map_wordToZ 0 "DEBUG" f
- | let G := get_goal in
- idtac "In the context:"; print_context ();
- idtac "In the goal:" G;
- idtac "When looking for bounds for" Zargs;
- check_is_map_wordToZ 1 "Expected a map of a function" f ];
- pose_proof_bounded_from_Zargs_hyps arg H
- | ?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 ?f ?x
- => let dummy := match goal with
- | _ => check_is_map_wordToZ 1 "In find_reified_f_evar, expected terms" f
- end in
- find_reified_f_evar x
- | _ => LHS
- end.
-Ltac zrange_to_reflective_hyps_step_gen then_tac round_up :=
- lazymatch 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 (@Bounds.bounds_to_base_type round_up) 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;
- then_tac ()
- | _ => idtac
- end.
-Ltac zrange_to_reflective_hyps_step := zrange_to_reflective_hyps_step_gen ltac:(fun _ => idtac).
-Ltac zrange_to_reflective_hyps round_up := zrange_to_reflective_hyps_step_gen ltac:(fun _ => zrange_to_reflective_hyps round_up) round_up.
-Ltac zrange_to_reflective_goal round_up 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 (@Bounds.bounds_to_base_type round_up) 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
- let mapped_output := constr:(map_output reified_f_evar) in
- let conjunct1 := constr:(is_bounded_by' output_bounds mapped_output) in
- let conjunct2_rhs := constr:(f (map_input args)) in
- let conjunct2 := constr:(mapped_output = conjunct2_rhs) in
- (* we use [cut] and [abstract] rather than [change] to catch
- inefficiencies in conversion early, rather than allowing
- [Defined] to take forever *)
- cut (conjunct1 /\ conjunct2);
- [ 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 round_up Hbounded :=
- zrange_to_reflective_hyps round_up; zrange_to_reflective_goal round_up Hbounded.
-
-
-(** ** [round_up_from_allowable_bit_widths] *)
-(** Construct a valid [round_up] function from allowable bit-widths *)
-Ltac round_up_from_allowable_bit_widths allowable_bit_widths :=
- let allowable_lgsz := (eval compute in (List.map Nat.log2 allowable_bit_widths)) in
- constr:(Bounds.round_up_to_in_list allowable_lgsz).
-
-(** ** [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' allowable_bit_widths Hbounded :=
- let round_up := round_up_from_allowable_bit_widths allowable_bit_widths in
- preunfold_and_dlet_to_context ();
- maybe_reassoc_sig_and_eexists ();
- intros_under_and ();
- do_curry_rhs ();
- split_BoundedWordToZ ();
- zrange_to_reflective round_up Hbounded.
-Ltac refine_to_reflective_glue allowable_bit_widths :=
- let Hbounded := fresh "Hbounded" in
- refine_to_reflective_glue' allowable_bit_widths Hbounded.
diff --git a/src/Compilers/Z/Bounds/Pipeline/OutputType.v b/src/Compilers/Z/Bounds/Pipeline/OutputType.v
deleted file mode 100644
index 028dd23a9..000000000
--- a/src/Compilers/Z/Bounds/Pipeline/OutputType.v
+++ /dev/null
@@ -1,52 +0,0 @@
-(** * 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.
-
-Section with_round_up_list.
- Context {allowable_lgsz : list nat}.
-
- Local Notation pick_typeb := (@Bounds.bounds_to_base_type (Bounds.round_up_to_in_list allowable_lgsz)) (only parsing).
- Local Notation pick_type v := (SmartFlatTypeMap pick_typeb v).
-
- Record ProcessedReflectivePackage
- := { InputType : _;
- input_expr : Expr 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 (Arrow (pick_type input_bounds) (pick_type output_bounds)) }.
-
- Definition Build_ProcessedReflectivePackage_from_option_sigma
- {t} (e : Expr 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 (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 InputType
- * { bounds : interp_flat_type Bounds.interp_base_type (domain InputType)
- * interp_flat_type Bounds.interp_base_type (codomain InputType)
- & Expr (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)).
-End with_round_up_list.
-
-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
deleted file mode 100644
index 9193ea17f..000000000
--- a/src/Compilers/Z/Bounds/Pipeline/ReflectiveTactics.v
+++ /dev/null
@@ -1,362 +0,0 @@
-(** * 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.Intros.
-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.InterpSideConditions.
-Require Import Crypto.Compilers.Z.InterpSideConditions.
-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.UnfoldArg.
-Require Import Crypto.Util.Tactics.UnifyAbstractReflexivity.
-Require Import Crypto.Util.FixedWordSizes.
-Require Import Crypto.Util.Option.
-Require Import bbv.WordScope.
-
-(** 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 :=
- unfold_second_arg Tuple.tuple;
- unfold_second_arg Tuple.tuple';
- 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_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.
-
-(** *** Boundedness Lemma Side Conditions *)
-(** The tactic [handle_boundedness_side_condition] unfolds relevant
- identifiers in the side-condition of the boundedness lemma. It
- then calls [boundedness_side_condition_solver], which can be
- overridden. *)
-(** The tactic [boundedness_side_condition_solver] attempts to solve
- the algebraic side conditions of the boundedness lemma; it leaves
- them over if it cannot solve them. *)
-Ltac boundedness_side_condition_solver :=
- repeat match goal with
- | [ |- _ /\ _ ] => apply conj
- | [ |- ?x = ?x ] => reflexivity
- end;
- try abstract ring.
-Ltac handle_boundedness_side_condition :=
- post_intro_interp_flat_type_intros;
- cbv [id
- fst snd
- InterpSideConditions Compilers.InterpSideConditions.InterpSideConditions interp_side_conditions interpf_side_conditions interpf_side_conditions_gen
- Z.Syntax.Util.interped_op_side_conditions
- ExprInversion.invert_Abs
- Syntax.interp_op Syntax.lift_op
- Syntax.Zinterp_op
- SmartMap.SmartFlatTypeMapUnInterp SmartMap.SmartValf SmartMap.SmartFlatTypeMapInterp2
- Syntax.cast_const Syntax.ZToInterp Syntax.interpToZ];
- cbv [PointedProp.and_pointed_Prop PointedProp.to_prop];
- try match goal with
- | [ |- True ] => exact I
- end;
- boundedness_side_condition_solver.
-
-(** ** 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.PointedProp.
-Require Import Crypto.Util.PartiallyReifiedProp.
-Require Import Crypto.Util.Equality.
-
-(** *** Gallina assembly *)
-Section with_round_up_list.
- Context {allowable_lgsz : list nat}.
-
- Local Notation pick_typeb := (@Bounds.bounds_to_base_type (Bounds.round_up_to_in_list allowable_lgsz)) (only parsing).
- Local Notation pick_type v := (SmartFlatTypeMap pick_typeb v).
-
- Context (opts : PipelineOptions)
- {t : type base_type}
- {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 interp_base_type (pick_type input_bounds)}
- {fZ : interp_flat_type interp_base_type (domain t)
- -> interp_flat_type interp_base_type (codomain t)}
- {final_e_evar : interp_flat_type interp_base_type (pick_type given_output_bounds)}.
-
- Class PipelineEvars :=
- {
- b : interp_flat_type Bounds.interp_base_type (codomain t);
- e' : Expr (pick_type input_bounds -> pick_type b);
- e_final_newtype : Expr (pick_type input_bounds -> pick_type given_output_bounds);
- e : Expr (domain t -> codomain t);
- e_pre_pkg : Expr (domain t -> codomain t);
- e_pkg : option (@ProcessedReflectivePackage allowable_lgsz)
- }.
-
- Context (evars : PipelineEvars)
-
- (** ** reification *)
- (rexpr_sig : { rexpr : Expr t | forall x, Interp rexpr x = fZ x }).
-
- Record PipelineSideConditions :=
- {
- (** ** 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-pre-bounds-pipeline *)
- Hpost_wf_pre_bounds : e_pre_pkg = PostWfPreBoundsPipeline opts e;
- (** ** post-wf-pipeline *)
- Hpost : e_pkg = PostWfBoundsPipeline _ opts e e_pre_pkg input_bounds;
- Hpost_correct : Some {| OutputType.input_expr := e ; OutputType.input_bounds := input_bounds ; output_bounds := b ; output_expr := e' |} = e_pkg;
- (** ** bounds relaxation *)
- Hbounds_relax : Bounds.is_tighter_thanb b given_output_bounds = true;
- Hbounds_sane : pick_type given_output_bounds = pick_type b;
- Hbounds_sane_refl
- : e_final_newtype
- = eq_rect _ (fun t => Expr (Arrow (pick_type input_bounds) t)) e' _ (eq_sym Hbounds_sane);
- (** ** instantiation of original evar *)
- Hevar : final_e_evar = InterpEta (t:=Arrow _ _) Syntax.interp_op e_final_newtype v';
- (** ** side conditions (boundedness) *)
- Hv1 : Bounds.is_bounded_by input_bounds (cast_back_flat_const v');
- Hv2 : intros_interp_flat_type_Prop (fun v => PointedProp.to_prop (InterpSideConditions e_pre_pkg v));
- }.
-
- Definition PipelineCorrect
- (side_conditions : PipelineSideConditions)
- : 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 using All.
- destruct side_conditions
- as [He He_unnatize_for_wf Hwf Hpost_wf_pre_bounds Hpost Hpost_correct Hbounds_relax Hbounds_sane Hbounds_sane_refl Hevar Hv1 Hv2].
- cbv [b e' e_final_newtype e e_pre_pkg e_pkg] in *.
- destruct evars as [b e' e_final_newtype e e_pre_pkg e_pkg];
- clear evars.
- destruct rexpr_sig as [? Hrexpr]; clear rexpr_sig.
- 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 eq_InterpEta, <- Hrexpr.
- pose proof (introsP_interp_flat_type Hv2) as Hv2'.
- 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 e v) k)
- end.
- rewrite <- transport_pp, concat_Vp; simpl.
- apply Hpost_correct.
- Qed.
-End with_round_up_list.
-
-(** ** 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. *)
-(** The [refine_PipelineSideConditions_constructor] subtactic splits
- up the [PipelineSideConditions] record goal into multiple
- subgoals, to be discharged by automation lower in this file
- ([solve_post_reified_side_conditions]). *)
-Ltac refine_PipelineSideConditions_constructor :=
- lazymatch goal with
- | [ |- PipelineSideConditions ?opts ?evars _ ]
- => simple refine {| Hevar := _ |};
- cbv [b e' e_final_newtype e e_pre_pkg e_pkg proj1_sig]
- end.
-Ltac refine_with_pipeline_correct opts :=
- lazymatch goal with
- | [ |- _ /\ ?castback ?fW = ?fZ ?arg ]
- => let lem := open_constr:(@PipelineCorrect _ opts _ _ _ _ _ _ {| e_pkg := _ |}) in
- simple refine (lem _ _);
- subst fW fZ
- end;
- [ eexists
- | refine_PipelineSideConditions_constructor ].
-
-(** 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_post_reified_side_conditions :=
- [>
- (** ** 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-pre-bounds-pipeline *)
- unify_abstract_vm_compute_rhs_reflexivity |
- (** ** post-wf pipeline *)
- unify_abstract_vm_compute_rhs_reflexivity |
- (** ** post-wf pipeline gives [Some _] *)
- unify_abstract_rhs_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) |
- (** ** 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 |
- (** ** 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 |
- (** ** boundedness side-condition *)
- handle_boundedness_side_condition ].
-Ltac solve_side_conditions :=
- [>
- (** ** reification *)
- do_reify |
- .. ];
- solve_post_reified_side_conditions.
-
-(** ** 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 opts :=
- refine_with_pipeline_correct opts; solve_side_conditions.
-Notation default_PipelineOptions := default_PipelineOptions.
diff --git a/src/Compilers/Z/Bounds/Relax.v b/src/Compilers/Z/Bounds/Relax.v
deleted file mode 100644
index 8178592ed..000000000
--- a/src/Compilers/Z/Bounds/Relax.v
+++ /dev/null
@@ -1,135 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.Arith.Arith.
-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.Compilers.Z.Bounds.RoundUpLemmas.
-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.Option.
-Require Import Crypto.Util.ZUtil.Log2.
-Require Import Crypto.Util.ZUtil.Tactics.LtbToLt.
-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'
- round_up
- t (tight_output_bounds relaxed_output_bounds : interp_flat_type Bounds.interp_base_type t)
- (Hv : SmartFlatTypeMap (@Bounds.bounds_to_base_type round_up) relaxed_output_bounds
- = SmartFlatTypeMap (@Bounds.bounds_to_base_type round_up) 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 (@Bounds.bounds_to_base_type round_up) tight_output_bounds
- v')
- /\ @cast_back_flat_const
- (@Bounds.interp_base_type) t (@Bounds.bounds_to_base_type round_up) 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 (@Bounds.bounds_to_base_type round_up) relaxed_output_bounds
- v)
- /\ @cast_back_flat_const
- (@Bounds.interp_base_type) t (@Bounds.bounds_to_base_type round_up) 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' Bounds.bounds_to_base_type ZRange.is_tighter_than_bool is_true SmartFlatTypeMap ZRange.is_bounded_by' Bounds.smallest_logsz 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_constr
- | progress subst
- | progress destruct_head bounds
- | progress destruct_head base_type
- | 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
- round_up
- t (tight_output_bounds relaxed_output_bounds : interp_flat_type Bounds.interp_base_type t)
- (Hv : SmartFlatTypeMap (@Bounds.bounds_to_base_type round_up) relaxed_output_bounds
- = SmartFlatTypeMap (@Bounds.bounds_to_base_type round_up) 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 (@Bounds.bounds_to_base_type round_up) 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 (@Bounds.bounds_to_base_type round_up) relaxed_output_bounds
- v
- = k.
-Proof.
- pose proof (fun pf => @relax_output_bounds' round_up 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/Bounds/RoundUpLemmas.v b/src/Compilers/Z/Bounds/RoundUpLemmas.v
deleted file mode 100644
index ce4ea9373..000000000
--- a/src/Compilers/Z/Bounds/RoundUpLemmas.v
+++ /dev/null
@@ -1,34 +0,0 @@
-Require Import Coq.omega.Omega.
-Require Import Crypto.Compilers.Z.Bounds.Interpretation.
-Require Import Crypto.Util.ZRange.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.NatUtil.
-
-Notation round_up_monotone_T round_up
- := (forall x y,
- (x <= y)%nat
- -> match round_up x, round_up y with
- | Some x', Some y' => (x' <= y')%nat
- | None, None => True
- | Some _, None => True
- | None, Some _ => False
- end)
- (only parsing).
-
-Lemma round_up_to_in_list_monotone (allowable_lgsz : list nat)
- : round_up_monotone_T (Bounds.round_up_to_in_list allowable_lgsz).
-Proof.
- intros x y H.
- induction allowable_lgsz as [|s ss].
- { constructor. }
- { unfold Bounds.round_up_to_in_list in *.
- repeat match goal with
- | _ => solve [ trivial ]
- | _ => progress simpl in *
- | _ => progress break_innermost_match_step
- | _ => progress break_innermost_match_hyps_step
- | [ H : ?leb _ _ = true |- _ ] => apply NPeano.Nat.leb_le in H
- | [ H : ?leb _ _ = false |- _ ] => apply NPeano.Nat.leb_gt in H
- | _ => omega *
- end. }
-Qed.
diff --git a/src/Compilers/Z/CNotations.v b/src/Compilers/Z/CNotations.v
deleted file mode 100644
index 07e2031f3..000000000
--- a/src/Compilers/Z/CNotations.v
+++ /dev/null
@@ -1,12696 +0,0 @@
-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 "T x = A ; 'return' b" (at level 200, b at level 200, format "T x = A ; '//' 'return' b").
-Reserved Notation "T x = A ; 'return' ( b0 , b1 , .. , b2 )" (at level 200, format "T x = A ; '//' 'return' ( b0 , b1 , .. , b2 )").
-Reserved Notation "T0 x , T1 y = A ; b" (at level 200, b at level 200, format "T0 x , T1 y = A ; '//' b").
-Reserved Notation "T0 x , T1 y = A ; 'return' b" (at level 200, b at level 200, format "T0 x , T1 y = A ; '//' 'return' b").
-Reserved Notation "T0 x , T1 y = A ; 'return' ( b0 , b1 , .. , b2 )" (at level 200, format "T0 x , T1 y = A ; '//' 'return' ( b0 , b1 , .. , b2 )").
-Reserved Notation "v == 0 ? a : b" (at level 40, a at level 10, b at level 10).
-Reserved Notation "v == 0 ?ℤ a : b" (at level 40, a at level 10, b at level 10).
-Reserved Notation "x & y" (at level 40).
-Reserved Notation "x | y" (at level 45).
-
-Global Open Scope expr_scope.
-
-Notation "T x = A ; b" := (LetIn (tx:=T) A (fun x => b)) : expr_scope.
-Notation "T x = A ; 'return' b" := (LetIn (tx:=T) A (fun x => Var b)) : expr_scope.
-Notation "T x = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=T) A (fun x => Pair .. (Pair b0%expr b1%expr) .. b2%expr)) : expr_scope.
-Notation "T x = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=T) A (fun x => Pair .. (Pair (Var b0) (Var b1)) .. (Var b2))) : expr_scope.
-Notation "T0 x , T1 y = A ; b" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => b)) : expr_scope.
-Notation "T0 x , T1 y = A ; 'return' b" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => Var b)) : expr_scope.
-(*Notation "T0 x , T1 y = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => (Pair .. (Pair b0%expr b1%expr) .. b2%expr))) : expr_scope.*) (* Error: Unsupported construction in recursive notations., https://coq.inria.fr/bugs/show_bug.cgi?id=5523 *)
-(*Notation "T0 x , T1 y = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => (Pair .. (Pair (Var b0) (Var b1)) .. (Var b2)))) : expr_scope.*) (* Error: Unsupported construction in recursive notations., https://coq.inria.fr/bugs/show_bug.cgi?id=5523 *)
-
-(* for now, handle with
-<<
-sed s':^\([^,]*\) \([^, ]*\)\(\s*\),\(.*\)\(addcarryx.*\))\([; ]*\)$:\1 \2\3;\4_\5, \&\2)\6:'
-sed s':^\([^,]*\) \([^, ]*\)\(\s*\),\(.*\)\(subborrow.*\))\([; ]*\)$:\1 \2\3;\4_\5, \&\2)\6:'
-sed s':^\([^,]*\) \([^, ]*\)\(\s*\),\(.*\)\(mulx.*\))\([; ]*\)$:\1 \2\3;\4_\5, \&\2)\6:'
->>
-
- Once we get https://coq.inria.fr/bugs/show_bug.cgi?id=5526, we can print actual C notations:
-<<
-Reserved Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c_in , a , b , & out ) ; REST"
- (at level 200, REST at level 200, only printing format "T0 out ; '//' T1 c_out = '_addcarryx_u32' ( c_in , a , b , & out ) ; '//' REST").
-Reserved Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c_in , a , b , & out ) ; REST"
- (at level 200, REST at level 200, only printing format "T0 out ; '//' T1 c_out = '_addcarryx_u64' ( c_in , a , b , & out ) ; '//' REST").
->> *)
-Reserved Notation "'addcarryx_u32' ( c , a , b )" (format "'addcarryx_u32' ( c , a , b )").
-Reserved Notation "'addcarryx_u64' ( c , a , b )" (format "'addcarryx_u64' ( c , a , b )").
-Reserved Notation "'addcarryx_u128' ( c , a , b )" (format "'addcarryx_u128' ( c , a , b )").
-Reserved Notation "'addcarryx_u25' ( c , a , b )" (format "'addcarryx_u25' ( c , a , b )").
-Reserved Notation "'addcarryx_u26' ( c , a , b )" (format "'addcarryx_u26' ( c , a , b )").
-Reserved Notation "'addcarryx_u51' ( c , a , b )" (format "'addcarryx_u51' ( c , a , b )").
-Reserved Notation "'subborrow_u32' ( c , a , b )" (format "'subborrow_u32' ( c , a , b )").
-Reserved Notation "'subborrow_u64' ( c , a , b )" (format "'subborrow_u64' ( c , a , b )").
-Reserved Notation "'subborrow_u128' ( c , a , b )" (format "'subborrow_u128' ( c , a , b )").
-Reserved Notation "'subborrow_u25' ( c , a , b )" (format "'subborrow_u25' ( c , a , b )").
-Reserved Notation "'subborrow_u26' ( c , a , b )" (format "'subborrow_u26' ( c , a , b )").
-Reserved Notation "'subborrow_u51' ( c , a , b )" (format "'subborrow_u51' ( c , a , b )").
-Reserved Notation "'addcarryx_u32ℤ' ( c , a , b )" (format "'addcarryx_u32ℤ' ( c , a , b )").
-Reserved Notation "'addcarryx_u64ℤ' ( c , a , b )" (format "'addcarryx_u64ℤ' ( c , a , b )").
-Reserved Notation "'addcarryx_u128ℤ' ( c , a , b )" (format "'addcarryx_u128ℤ' ( c , a , b )").
-Reserved Notation "'addcarryx_u25ℤ' ( c , a , b )" (format "'addcarryx_u25ℤ' ( c , a , b )").
-Reserved Notation "'addcarryx_u26ℤ' ( c , a , b )" (format "'addcarryx_u26ℤ' ( c , a , b )").
-Reserved Notation "'addcarryx_u51ℤ' ( c , a , b )" (format "'addcarryx_u51ℤ' ( c , a , b )").
-Reserved Notation "'subborrow_u32ℤ' ( c , a , b )" (format "'subborrow_u32ℤ' ( c , a , b )").
-Reserved Notation "'subborrow_u64ℤ' ( c , a , b )" (format "'subborrow_u64ℤ' ( c , a , b )").
-Reserved Notation "'subborrow_u128ℤ' ( c , a , b )" (format "'subborrow_u128ℤ' ( c , a , b )").
-Reserved Notation "'subborrow_u25ℤ' ( c , a , b )" (format "'subborrow_u25ℤ' ( c , a , b )").
-Reserved Notation "'subborrow_u26ℤ' ( c , a , b )" (format "'subborrow_u26ℤ' ( c , a , b )").
-Reserved Notation "'subborrow_u51ℤ' ( c , a , b )" (format "'subborrow_u51ℤ' ( c , a , b )").
-
-Reserved Notation "'mulx_u32' ( a , b )" (format "'mulx_u32' ( a , b )").
-Reserved Notation "'mulx_u64' ( a , b )" (format "'mulx_u64' ( a , b )").
-Reserved Notation "'mulx_u128' ( a , b )" (format "'mulx_u128' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u32' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u64' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u128' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u32' ( a , b )" (format "'(uint8_t)' 'mulx_u32' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u64' ( a , b )" (format "'(uint8_t)' 'mulx_u64' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u128' ( a , b )" (format "'(uint8_t)' 'mulx_u128' ( a , b )").
-Reserved Notation "'mulx_u32_out_u8' ( a , b )" (format "'mulx_u32_out_u8' ( a , b )").
-Reserved Notation "'mulx_u64_out_u8' ( a , b )" (format "'mulx_u64_out_u8' ( a , b )").
-Reserved Notation "'mulx_u128_out_u8' ( a , b )" (format "'mulx_u128_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u32_out_u8' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u32_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u64_out_u8' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u64_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u128_out_u8' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u128_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" (format "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" (format "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" (format "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )").
-Reserved Notation "'mulx_u32_out_u1' ( a , b )" (format "'mulx_u32_out_u1' ( a , b )").
-Reserved Notation "'mulx_u64_out_u1' ( a , b )" (format "'mulx_u64_out_u1' ( a , b )").
-Reserved Notation "'mulx_u128_out_u1' ( a , b )" (format "'mulx_u128_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u32_out_u1' ( a , b )" (format "'(uint8_t)' 'mulx_u32_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u64_out_u1' ( a , b )" (format "'(uint8_t)' 'mulx_u64_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u128_out_u1' ( a , b )" (format "'(uint8_t)' 'mulx_u128_out_u1' ( a , b )").
-
-Reserved Notation "'mulx_u32ℤ' ( a , b )" (format "'mulx_u32ℤ' ( a , b )").
-Reserved Notation "'mulx_u64ℤ' ( a , b )" (format "'mulx_u64ℤ' ( a , b )").
-Reserved Notation "'mulx_u128ℤ' ( a , b )" (format "'mulx_u128ℤ' ( a , b )").
-
-Reserved Notation "'cmovznz32' ( v , a , b )" (format "'cmovznz32' ( v , a , b )").
-Reserved Notation "'cmovznz64' ( v , a , b )" (format "'cmovznz64' ( v , a , b )").
-Reserved Notation "'cmovznz128' ( v , a , b )" (format "'cmovznz128' ( v , a , b )").
-Reserved Notation "'cmovznz' ( v , a , b )" (format "'cmovznz' ( v , a , b )").
-Reserved Notation "'cmovznzℤ' ( v , a , b )" (format "'cmovznzℤ' ( v , a , b )").
-(* python:
-<<
-#!/usr/bin/env python
-# -*- coding: utf-8 -*-
-import math
-
-PARENTHESIZED = True
-ADD_CARRY_SUB_BORROW_SIZES = (32, 64, 128, 25, 26, 51)
-
-print(r"""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 "T x = A ; 'return' b" (at level 200, b at level 200, format "T x = A ; '//' 'return' b").
-Reserved Notation "T x = A ; 'return' ( b0 , b1 , .. , b2 )" (at level 200, format "T x = A ; '//' 'return' ( b0 , b1 , .. , b2 )").
-Reserved Notation "T0 x , T1 y = A ; b" (at level 200, b at level 200, format "T0 x , T1 y = A ; '//' b").
-Reserved Notation "T0 x , T1 y = A ; 'return' b" (at level 200, b at level 200, format "T0 x , T1 y = A ; '//' 'return' b").
-Reserved Notation "T0 x , T1 y = A ; 'return' ( b0 , b1 , .. , b2 )" (at level 200, format "T0 x , T1 y = A ; '//' 'return' ( b0 , b1 , .. , b2 )").
-Reserved Notation "v == 0 ? a : b" (at level 40, a at level 10, b at level 10).
-Reserved Notation "v == 0 ?ℤ a : b" (at level 40, a at level 10, b at level 10).
-Reserved Notation "x & y" (at level 40).
-Reserved Notation "x | y" (at level 45).
-
-Global Open Scope expr_scope.
-
-Notation "T x = A ; b" := (LetIn (tx:=T) A (fun x => b)) : expr_scope.
-Notation "T x = A ; 'return' b" := (LetIn (tx:=T) A (fun x => Var b)) : expr_scope.
-Notation "T x = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=T) A (fun x => Pair .. (Pair b0%expr b1%expr) .. b2%expr)) : expr_scope.
-Notation "T x = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=T) A (fun x => Pair .. (Pair (Var b0) (Var b1)) .. (Var b2))) : expr_scope.
-Notation "T0 x , T1 y = A ; b" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => b)) : expr_scope.
-Notation "T0 x , T1 y = A ; 'return' b" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => Var b)) : expr_scope.
-(""" + r"""*Notation "T0 x , T1 y = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => (Pair .. (Pair b0%expr b1%expr) .. b2%expr))) : expr_scope.*""" + r""") (""" + r"""* Error: Unsupported construction in recursive notations., https://coq.inria.fr/bugs/show_bug.cgi?id=5523 *""" + r""")
-(""" + r"""*Notation "T0 x , T1 y = A ; 'return' ( b0 , b1 , .. , b2 )" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => (Pair .. (Pair (Var b0) (Var b1)) .. (Var b2)))) : expr_scope.*""" + r""") (""" + r"""* Error: Unsupported construction in recursive notations., https://coq.inria.fr/bugs/show_bug.cgi?id=5523 *""" + r""")
-
-(""" + r"""* for now, handle with
-<<
-sed s':^\([^,]*\) \([^, ]*\)\(\s*\),\(.*\)\(addcarryx.*\))\([; ]*\)$:\1 \2\3;\4_\5, \&\2)\6:'
-sed s':^\([^,]*\) \([^, ]*\)\(\s*\),\(.*\)\(subborrow.*\))\([; ]*\)$:\1 \2\3;\4_\5, \&\2)\6:'
-sed s':^\([^,]*\) \([^, ]*\)\(\s*\),\(.*\)\(mulx.*\))\([; ]*\)$:\1 \2\3;\4_\5, \&\2)\6:'
->>
-
- Once we get https://coq.inria.fr/bugs/show_bug.cgi?id=5526, we can print actual C notations:
-<<
-Reserved Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c_in , a , b , & out ) ; REST"
- (at level 200, REST at level 200, only printing format "T0 out ; '//' T1 c_out = '_addcarryx_u32' ( c_in , a , b , & out ) ; '//' REST").
-Reserved Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c_in , a , b , & out ) ; REST"
- (at level 200, REST at level 200, only printing format "T0 out ; '//' T1 c_out = '_addcarryx_u64' ( c_in , a , b , & out ) ; '//' REST").
->> *""" + r""")""")
-for postfix in ('', 'ℤ'):
- for opn in ('addcarryx', 'subborrow'):
- for sz in ADD_CARRY_SUB_BORROW_SIZES:
- print(r"""Reserved Notation "'%s_u%d%s' ( c , a , b )" (format "'%s_u%d%s' ( c , a , b )").""" % (opn, sz, postfix, opn, sz, postfix))
-print(r"""
-Reserved Notation "'mulx_u32' ( a , b )" (format "'mulx_u32' ( a , b )").
-Reserved Notation "'mulx_u64' ( a , b )" (format "'mulx_u64' ( a , b )").
-Reserved Notation "'mulx_u128' ( a , b )" (format "'mulx_u128' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u32' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u64' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u128' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u32' ( a , b )" (format "'(uint8_t)' 'mulx_u32' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u64' ( a , b )" (format "'(uint8_t)' 'mulx_u64' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u128' ( a , b )" (format "'(uint8_t)' 'mulx_u128' ( a , b )").
-Reserved Notation "'mulx_u32_out_u8' ( a , b )" (format "'mulx_u32_out_u8' ( a , b )").
-Reserved Notation "'mulx_u64_out_u8' ( a , b )" (format "'mulx_u64_out_u8' ( a , b )").
-Reserved Notation "'mulx_u128_out_u8' ( a , b )" (format "'mulx_u128_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u32_out_u8' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u32_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u64_out_u8' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u64_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u128_out_u8' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u128_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" (format "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" (format "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" (format "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )").
-Reserved Notation "'mulx_u32_out_u1' ( a , b )" (format "'mulx_u32_out_u1' ( a , b )").
-Reserved Notation "'mulx_u64_out_u1' ( a , b )" (format "'mulx_u64_out_u1' ( a , b )").
-Reserved Notation "'mulx_u128_out_u1' ( a , b )" (format "'mulx_u128_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" (format "'(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u32_out_u1' ( a , b )" (format "'(uint8_t)' 'mulx_u32_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u64_out_u1' ( a , b )" (format "'(uint8_t)' 'mulx_u64_out_u1' ( a , b )").
-Reserved Notation "'(uint8_t)' 'mulx_u128_out_u1' ( a , b )" (format "'(uint8_t)' 'mulx_u128_out_u1' ( a , b )").
-
-Reserved Notation "'mulx_u32ℤ' ( a , b )" (format "'mulx_u32ℤ' ( a , b )").
-Reserved Notation "'mulx_u64ℤ' ( a , b )" (format "'mulx_u64ℤ' ( a , b )").
-Reserved Notation "'mulx_u128ℤ' ( a , b )" (format "'mulx_u128ℤ' ( a , b )").
-
-Reserved Notation "'cmovznz32' ( v , a , b )" (format "'cmovznz32' ( v , a , b )").
-Reserved Notation "'cmovznz64' ( v , a , b )" (format "'cmovznz64' ( v , a , b )").
-Reserved Notation "'cmovznz128' ( v , a , b )" (format "'cmovznz128' ( v , a , b )").
-Reserved Notation "'cmovznz' ( v , a , b )" (format "'cmovznz' ( v , a , b )").
-Reserved Notation "'cmovznzℤ' ( v , a , b )" (format "'cmovznzℤ' ( v , a , b )").
-(""" + r"""* python:
-<<""")
-
-print(open(__file__).read())
-
-print('>> *' + ')')
-
-OPEN = ('( ' if PARENTHESIZED else '')
-CLOSE = (' )' if PARENTHESIZED else '')
-
-def log2_up(x):
- return int(math.ceil(math.log(x, 2)))
-types = ('uint8_t/*bool*/', 'uint8_t', 'uint8_t', 'uint8_t', 'uint16_t', 'uint32_t', 'uint64_t', 'uint128_t', 'uint256_t')
-for lgwordsz in range(0, len(types)):
- print('Notation "\'%s\'" := (Tbase (TWord %d)) : expr_scope.' % (types[lgwordsz], lgwordsz))
-print('Notation ℤ := (Tbase TZ).')
-print('')
-cast_pat = "%s %s"
-cast_types = tuple(["'1&(%s)'" % i for i in types[:1]] +
- ["'(%s)'" % i for i in types[1:]])
-def at_least_S_pattern(n):
- return '(S ' * n + '_' + ')' * n
-
-_ = '_'
-TZ = 'TZ'
-def TWord(v): return '(TWord %s)' % str(v)
-
-def print_notation_string(xisvar, yisvar, opn, op, arg_tuple, lvl=None, xsz=None, ysz=None, xlvl=None, ylvl=None):
- lhs = ('x' if not xisvar else '(Var x)')
- rhs = ('y' if not yisvar else '(Var y)')
- x = ('x' if xsz is None else (cast_pat % (cast_types[xsz], 'x')))
- y = ('y' if ysz is None else (cast_pat % (cast_types[ysz], 'y')))
- ret = 'Notation "%s%s %s %s%s" := (Op (%s %s) (Pair %s %s))' % (OPEN, x, opn, y, CLOSE, op, ' '.join(arg_tuple), lhs, rhs)
- modifiers = []
- for var, l in (('', lvl), ('x ', xlvl), ('y ', ylvl)):
- if l is not None:
- modifiers.append('%sat level %s' % (var, str(l)))
- if (OPEN + CLOSE) != '':
- modifiers.append('format "%s%s %s %s%s"' % (OPEN, x, opn, y, CLOSE))
- if len(modifiers) > 0:
- ret = '%s (%s) : expr_scope.' % (ret, ', '.join(modifiers))
- else:
- ret = '%s : expr_scope.' % (ret,)
- print(ret)
-
-for opn, op, lvl in (('*', 'Mul', 40), ('+', 'Add', 50), ('-', 'Sub', 50)):
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (_, _, _))
- print_notation_string(v1, v2, opn + 'ℤ', op, (_, _, TZ), lvl=lvl)
- for lgwordsz in range(0, len(types)):
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (TWord(_), TWord(_), TWord(lgwordsz)), xsz=lgwordsz, lvl=lvl, xlvl=9)
- print_notation_string(v1, v2, opn, op, (TWord(_), TWord(at_least_S_pattern(lgwordsz)), TWord(lgwordsz)), ysz=lgwordsz, lvl=lvl, ylvl=9)
- print_notation_string(v1, v2, opn, op, (TWord(at_least_S_pattern(lgwordsz)), TWord(at_least_S_pattern(lgwordsz)), TWord(lgwordsz)), xsz=lgwordsz, ysz=lgwordsz, lvl=lvl, xlvl=9, ylvl=9)
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (TWord(lgwordsz), TWord(_), TWord(lgwordsz)))
- print_notation_string(v1, v2, opn, op, (TWord(_), TWord(lgwordsz), TWord(lgwordsz)))
- print_notation_string(v1, v2, opn, op, (TWord(lgwordsz), TWord(at_least_S_pattern(lgwordsz)), TWord(lgwordsz)), ysz=lgwordsz, lvl=lvl, ylvl=9)
- print_notation_string(v1, v2, opn, op, (TWord(at_least_S_pattern(lgwordsz)), TWord(lgwordsz), TWord(lgwordsz)), xsz=lgwordsz, lvl=lvl, xlvl=9)
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (TWord(lgwordsz), TWord(lgwordsz), TWord(lgwordsz)))
-for opn, op, lvl in (('*', 'Mul', 40), ('+', 'Add', 50), ('-', 'Sub', 50)):
- for v1 in (False, True):
- for v2 in (False, True):
- x = ('x' if not v1 else '(Var x)')
- y = ('y' if not v2 else '(Var y)')
- print('''Notation "%s'1&(' %s %s %s ')'%s" := (Op (%s %s) (Pair %s %s)) (at level %d, format "%s'1&(' %s %s %s ')'%s").''' %
- (OPEN, 'x', opn, 'y', CLOSE,
- op, ' '.join((TWord(_), TWord(_), TWord(0))), x, y, lvl,
- OPEN, 'x', opn, 'y', CLOSE))
-for opn, op, lvl in (('&', 'Land', 40), ('|', 'Lor', 45)):
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (_, _, _))
- print_notation_string(v1, v2, opn + 'ℤ', op, (_, _, TZ), lvl=lvl)
- for lgwordsz in range(0, len(types)):
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (TWord(_), TWord(_), TWord(lgwordsz)), xsz=lgwordsz, ysz=lgwordsz, lvl=lvl, xlvl=9, ylvl=9)
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (TWord(lgwordsz), TWord(_), TWord(lgwordsz)), ysz=lgwordsz, lvl=lvl, ylvl=9)
- print_notation_string(v1, v2, opn, op, (TWord(_), TWord(lgwordsz), TWord(lgwordsz)), xsz=lgwordsz, lvl=lvl, xlvl=9)
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (TWord(lgwordsz), TWord(lgwordsz), TWord(lgwordsz)))
-for opn, op, lvl in (('<<', 'Shl', 30),):
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (_, _, _))
- print_notation_string(v1, v2, opn + 'ℤ', op, (_, _, TZ), lvl=lvl)
- for lgwordsz in range(0, len(types)):
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (TWord(_), TWord(_), TWord(lgwordsz)), xsz=lgwordsz, lvl=lvl)
- print_notation_string(v1, v2, opn, op, (TWord(lgwordsz), TWord(_), TWord(lgwordsz)))
-for opn, op, lvl in (('>>', 'Shr', 30),):
- for v1 in (False, True):
- for v2 in (False, True):
- print_notation_string(v1, v2, opn, op, (_, _, _))
- print_notation_string(v1, v2, opn + 'ℤ', op, (_, _, TZ), lvl=lvl)
- 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) : expr_scope.'
- % (types[lgwordsz], opn, op, lgwordsz, lhs, rhs, lvl))
- print_notation_string(v1, v2, opn, op, (TWord(lgwordsz), TWord(_), TWord(lgwordsz)))
-for v0 in (False, True):
- for v1 in (False, True):
- for v2 in (False, True):
- tes = ('v' if not v0 else '(Var v)')
- lhs = ('x' if not v1 else '(Var x)')
- rhs = ('y' if not v2 else '(Var y)')
- print('Notation "\'cmovznz32\' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair %s %s) %s)) : expr_scope.' % (tes, lhs, rhs))
-for (wordsz, cmovsz) in ((16,'32'), (32,'64'), (64,'128'), (128,'ℤ')):
- lgwordsz = log2_up(wordsz)
- for lgwordsz_out in range(0, lgwordsz+2): # + 2 so that we get one instance of lgwordsz_out > lgwordsz
- for v0 in (False, True):
- for v1 in (False, True):
- for v2 in (False, True):
- tes = ('v' if not v0 else '(Var v)')
- lhs = ('x' if not v1 else '(Var x)')
- rhs = ('y' if not v2 else '(Var y)')
- TWordSz = TWord(at_least_S_pattern(lgwordsz+1)) # + 1, because we need strictly not equal
- TWordSz_out = TWord(at_least_S_pattern(lgwordsz_out))
- if lgwordsz_out > lgwordsz:
- pat = 'Notation "\'cmovznz%s\' ( v , x , y )" := (Op (Zselect %s %s %s %s) (Pair (Pair %s %s) %s)) : expr_scope.'
- print(pat % (cmovsz, TWordSz, TWord(_), TWord(_), TWordSz_out, tes, lhs, rhs))
- print(pat % (cmovsz, TWord(_), TWordSz, TWord(_), TWordSz_out, tes, lhs, rhs))
- print(pat % (cmovsz, TWord(_), TWord(_), TWordSz, TWordSz_out, tes, lhs, rhs))
- else:
- pat = 'Notation "\'(%s)\' \'cmovznz%s\' ( v , x , y )" := (Op (Zselect %s %s %s %s) (Pair (Pair %s %s) %s)) (format "\'(%s)\' \'cmovznz%s\' ( v , x , y )") : expr_scope.'
- print(pat % (types[lgwordsz_out], cmovsz, TWordSz, TWord(_), TWord(_), TWordSz_out, tes, lhs, rhs, types[lgwordsz_out], cmovsz))
- print(pat % (types[lgwordsz_out], cmovsz, TWord(_), TWordSz, TWord(_), TWordSz_out, tes, lhs, rhs, types[lgwordsz_out], cmovsz))
- print(pat % (types[lgwordsz_out], cmovsz, TWord(_), TWord(_), TWordSz, TWordSz_out, tes, lhs, rhs, types[lgwordsz_out], cmovsz))
-for v0 in (False, True):
- for v1 in (False, True):
- for v2 in (False, True):
- tes = ('v' if not v0 else '(Var v)')
- lhs = ('x' if not v1 else '(Var x)')
- rhs = ('y' if not v2 else '(Var y)')
- print('Notation "\'cmovznzℤ\' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair %s %s) %s)) : expr_scope.' % (tes, lhs, rhs))
- print('Notation "\'cmovznzℤ\' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair %s %s) %s)) : expr_scope.' % (tes, lhs, rhs))
- print('Notation "\'cmovznzℤ\' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair %s %s) %s)) : expr_scope.' % (tes, lhs, rhs))
- print('Notation "\'cmovznzℤ\' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair %s %s) %s)) : expr_scope.' % (tes, lhs, rhs))
-for opn, op in (('addcarryx', 'AddWithGetCarry'), ('subborrow', 'SubWithGetBorrow')):
- for wordsz in ADD_CARRY_SUB_BORROW_SIZES:
- lgwordsz = log2_up(wordsz)
- for v0 in (False, True):
- for v1 in (False, True):
- for v2 in (False, True):
- c = ('c' if not v0 else '(Var c)')
- a = ('a' if not v1 else '(Var a)')
- b = ('b' if not v2 else '(Var b)')
- for lgwordsz_small in (0, ): # (0, 3):
- for a_lgwordsz in range(lgwordsz+1):
- for b_lgwordsz in range(lgwordsz+1):
- for notation_string in ('Notation "\'%s_u%d\' ( c , a , b )" := (Op (%s %d (TWord %d) (TWord %d) (TWord %d) (TWord %d) (TWord %d)) (Pair (Pair %s %s) %s)) : expr_scope.',
- ('(' + '*Notation "T0 out ; T1 c_out = \'_%s_u%d\' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (%s %d (TWord %d) (TWord %d) (TWord %d) (TWord %d) (TWord %d)) (Pair (Pair %s %s) %s)) (fun \'((out, c_out)%%core) => REST)) : expr_scope.*' + ')')):
- print(notation_string % (opn, wordsz, op, wordsz, lgwordsz_small, a_lgwordsz, b_lgwordsz, lgwordsz, lgwordsz_small, c, a, b))
-for opn, op in (('mulx', 'MulSplit'),):
- for wordsz in (32, 64, 128, 51):
- lgwordsz = log2_up(wordsz)
- for v0 in (False, True):
- for v1 in (False, True):
- a = ('a' if not v0 else '(Var a)')
- b = ('b' if not v1 else '(Var b)')
- for lgwordsz_small in (0, 3):
- for notation_string in ('Notation "%s\'%s_u%d%s\' ( a , b )" := (Op (%s %d (TWord %d) (TWord %d) (TWord %d) (TWord %d)) (Pair %s %s)) : expr_scope.',
- ('(' + '*Notation "T0 out ; T1 c_out = %s\'_%s_u%d%s\' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (%s %d (TWord %d) (TWord %d) (TWord %d) (TWord %d)) (Pair %s %s)) (fun \'((out, c_out)%%core) => REST)) : expr_scope.*' + ')')):
- for arg1 in (lgwordsz_small, lgwordsz):
- for arg2 in (lgwordsz_small, lgwordsz):
- for arg3 in (lgwordsz_small, lgwordsz):
- for arg4 in (lgwordsz_small, lgwordsz): # N.B. the final argument, which is the high bits, must be of a compatible pointer type, and cannot be a pointer to a short word, without invoking a separate function
- cast_val = ('' if arg3 == lgwordsz else (cast_pat % (cast_types[lgwordsz_small], '')))
- extra_fun = ('' if arg4 == lgwordsz else ('_out_u%d' % (2**arg4)))
- print(notation_string % (cast_val, opn, wordsz, extra_fun, op, wordsz, arg1, arg2, arg3, arg4, a, b))
-for opn, op in (('mulx', 'MulSplit'),):
- for wordsz in (32, 64, 128, 51):
- lgwordsz = log2_up(wordsz)
- for v0 in (False, True):
- for v1 in (False, True):
- a = ('a' if not v0 else '(Var a)')
- b = ('b' if not v1 else '(Var b)')
- print(('Notation "\'%s_u%d\' ( a , b )" := (Op (%s %d (TWord %d) (TWord %d) (TWord %d) (TWord %d)) (Pair %s %s)) : expr_scope.') % (opn, wordsz, op, wordsz, lgwordsz, lgwordsz, lgwordsz, lgwordsz, a, b))
- print(('(' + '*Notation "T0 out ; T1 c_out = \'_%s_u%d\' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (%s %d (TWord %d) (TWord %d) (TWord %d) (TWord %d)) (Pair %s %s)) (fun \'((out, c_out)%%core) => REST)) : expr_scope.*' + ')') % (opn, wordsz, op, wordsz, lgwordsz, lgwordsz, lgwordsz, lgwordsz, a, b))
-for opn, op in (('addcarryx', 'AddWithGetCarry'), ('subborrow', 'SubWithGetBorrow')):
- for wordsz in ADD_CARRY_SUB_BORROW_SIZES:
- lgwordsz = log2_up(wordsz)
- for v0 in (False, True):
- for v1 in (False, True):
- for v2 in (False, True):
- c = ('c' if not v0 else '(Var c)')
- a = ('a' if not v1 else '(Var a)')
- b = ('b' if not v2 else '(Var b)')
- print(('Notation "\'%s_u%dℤ\' ( c , a , b )" := (Op (%s %d _ _ _ _ TZ) (Pair (Pair %s %s) %s)) : expr_scope.') % (opn, wordsz, op, wordsz, c, a, b))
- print(('Notation "\'%s_u%dℤ\' ( c , a , b )" := (Op (%s %d _ _ _ TZ _) (Pair (Pair %s %s) %s)) : expr_scope.') % (opn, wordsz, op, wordsz, c, a, b))
- print(('(' + '*Notation "T0 out ; T1 c_out = \'_%s_u%dℤ\' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (%s %d _ _ _ _ TZ) (Pair (Pair %s %s) %s)) (fun \'((out, c_out)%%core) => REST)) : expr_scope.*' + ')') % (opn, wordsz, op, wordsz, c, a, b))
- print(('(' + '*Notation "T0 out ; T1 c_out = \'_%s_u%dℤ\' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (%s %d _ _ _ TZ _) (Pair (Pair %s %s) %s)) (fun \'((out, c_out)%%core) => REST)) : expr_scope.*' + ')') % (opn, wordsz, op, wordsz, c, a, b))
-for opn, op in (('mulx', 'MulSplit'),):
- for wordsz in (32, 64, 128, 51):
- lgwordsz = log2_up(wordsz)
- for v0 in (False, True):
- for v1 in (False, True):
- a = ('a' if not v0 else '(Var a)')
- b = ('b' if not v1 else '(Var b)')
- print(('Notation "\'%s_u%dℤ\' ( a , b )" := (Op (%s %d _ _ _ TZ) (Pair %s %s)) : expr_scope.') % (opn, wordsz, op, wordsz, a, b))
- print(('Notation "\'%s_u%dℤ\' ( a , b )" := (Op (%s %d _ _ TZ _) (Pair %s %s)) : expr_scope.') % (opn, wordsz, op, wordsz, a, b))
- print(('(' + '*Notation "T0 out ; T1 c_out = \'_%s_u%dℤ\' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (%s %d _ _ _ TZ) (Pair %s %s)) (fun \'((out, c_out)%%core) => REST)) : expr_scope.*' + ')') % (opn, wordsz, op, wordsz, a, b))
- print(('(' + '*Notation "T0 out ; T1 c_out = \'_%s_u%dℤ\' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (%s %d _ _ TZ _) (Pair %s %s)) (fun \'((out, c_out)%%core) => REST)) : expr_scope.*' + ')') % (opn, wordsz, op, wordsz, a, b))
-print('Notation Return x := (Var x).')
-print('Notation C_like := (Expr base_type op _).')
-
->> *)
-Notation "'uint8_t/*bool*/'" := (Tbase (TWord 0)) : expr_scope.
-Notation "'uint8_t'" := (Tbase (TWord 1)) : expr_scope.
-Notation "'uint8_t'" := (Tbase (TWord 2)) : expr_scope.
-Notation "'uint8_t'" := (Tbase (TWord 3)) : expr_scope.
-Notation "'uint16_t'" := (Tbase (TWord 4)) : expr_scope.
-Notation "'uint32_t'" := (Tbase (TWord 5)) : expr_scope.
-Notation "'uint64_t'" := (Tbase (TWord 6)) : expr_scope.
-Notation "'uint128_t'" := (Tbase (TWord 7)) : expr_scope.
-Notation "'uint256_t'" := (Tbase (TWord 8)) : expr_scope.
-Notation ℤ := (Tbase TZ).
-
-Notation "( x * y )" := (Op (Mul _ _ _) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x *ℤ y )" := (Op (Mul _ _ TZ) (Pair x y)) (at level 40, format "( x *ℤ y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul _ _ _) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x *ℤ y )" := (Op (Mul _ _ TZ) (Pair x (Var y))) (at level 40, format "( x *ℤ y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul _ _ _) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x *ℤ y )" := (Op (Mul _ _ TZ) (Pair (Var x) y)) (at level 40, format "( x *ℤ y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul _ _ _) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x *ℤ y )" := (Op (Mul _ _ TZ) (Pair (Var x) (Var y))) (at level 40, format "( x *ℤ y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '1&(uint8_t/*bool*/)' y )" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x * y )" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 1) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S _)) (TWord 1) (TWord 1)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 1) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S _)) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 1) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S _)) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 1) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S _)) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * '(uint8_t)' y )" := (Op (Mul (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x * '(uint8_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint8_t)' y )" := (Op (Mul (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x * y )" := (Op (Mul (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * '(uint16_t)' y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * '(uint16_t)' y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * '(uint16_t)' y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * '(uint16_t)' y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x * '(uint16_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair x y)) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint16_t)' y )" := (Op (Mul (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x * y )" := (Op (Mul (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint16_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * '(uint32_t)' y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * '(uint32_t)' y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * '(uint32_t)' y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * '(uint32_t)' y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x * '(uint32_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair x y)) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint32_t)' y )" := (Op (Mul (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint32_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * '(uint64_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * '(uint64_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * '(uint64_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * '(uint64_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x * '(uint64_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair x y)) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint64_t)' y )" := (Op (Mul (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint64_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * '(uint128_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * '(uint128_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * '(uint128_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * '(uint128_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x * '(uint128_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair x y)) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint128_t)' y )" := (Op (Mul (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint128_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * '(uint256_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * '(uint256_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * '(uint256_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * '(uint256_t)' y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x * '(uint256_t)' y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair x y)) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * '(uint256_t)' y )" := (Op (Mul (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x * '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x * y )" := (Op (Mul (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint256_t)' x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (Pair x y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x * y )") : expr_scope.
-Notation "( x * y )" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x * y )") : expr_scope.
-Notation "( x + y )" := (Op (Add _ _ _) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x +ℤ y )" := (Op (Add _ _ TZ) (Pair x y)) (at level 50, format "( x +ℤ y )") : expr_scope.
-Notation "( x + y )" := (Op (Add _ _ _) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x +ℤ y )" := (Op (Add _ _ TZ) (Pair x (Var y))) (at level 50, format "( x +ℤ y )") : expr_scope.
-Notation "( x + y )" := (Op (Add _ _ _) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x +ℤ y )" := (Op (Add _ _ TZ) (Pair (Var x) y)) (at level 50, format "( x +ℤ y )") : expr_scope.
-Notation "( x + y )" := (Op (Add _ _ _) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x +ℤ y )" := (Op (Add _ _ TZ) (Pair (Var x) (Var y))) (at level 50, format "( x +ℤ y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '1&(uint8_t/*bool*/)' y )" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x + y )" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 1) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S _)) (TWord 1) (TWord 1)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 1) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S _)) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 1) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S _)) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 1) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S _)) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + '(uint8_t)' y )" := (Op (Add (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x + '(uint8_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint8_t)' y )" := (Op (Add (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x + y )" := (Op (Add (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + '(uint16_t)' y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + '(uint16_t)' y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + '(uint16_t)' y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + '(uint16_t)' y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x + '(uint16_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair x y)) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint16_t)' y )" := (Op (Add (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x + y )" := (Op (Add (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + '(uint32_t)' y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + '(uint32_t)' y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + '(uint32_t)' y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + '(uint32_t)' y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x + '(uint32_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair x y)) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint32_t)' y )" := (Op (Add (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + '(uint64_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + '(uint64_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + '(uint64_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + '(uint64_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x + '(uint64_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair x y)) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint64_t)' y )" := (Op (Add (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + '(uint128_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + '(uint128_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + '(uint128_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + '(uint128_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x + '(uint128_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair x y)) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint128_t)' y )" := (Op (Add (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + '(uint256_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + '(uint256_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + '(uint256_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + '(uint256_t)' y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x + '(uint256_t)' y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair x y)) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + '(uint256_t)' y )" := (Op (Add (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x + '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x + y )" := (Op (Add (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (Pair x y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x + y )") : expr_scope.
-Notation "( x + y )" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x + y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub _ _ _) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x -ℤ y )" := (Op (Sub _ _ TZ) (Pair x y)) (at level 50, format "( x -ℤ y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub _ _ _) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x -ℤ y )" := (Op (Sub _ _ TZ) (Pair x (Var y))) (at level 50, format "( x -ℤ y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub _ _ _) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x -ℤ y )" := (Op (Sub _ _ TZ) (Pair (Var x) y)) (at level 50, format "( x -ℤ y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub _ _ _) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x -ℤ y )" := (Op (Sub _ _ TZ) (Pair (Var x) (Var y))) (at level 50, format "( x -ℤ y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '1&(uint8_t/*bool*/)' y )" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x - y )" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '1&(uint8_t/*bool*/)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S _)) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 1) (TWord (S _)) (TWord 1)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S _)) (TWord 1) (TWord 1)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 1) (TWord (S _)) (TWord 1)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S _)) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 1) (TWord (S _)) (TWord 1)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S _)) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 1) (TWord (S _)) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S _)) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S _))) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 2) (TWord (S (S _))) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S _))) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - '(uint8_t)' y )" := (Op (Sub (TWord (S (S (S _)))) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint8_t)' x - '(uint8_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair x y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint8_t)' y )" := (Op (Sub (TWord 3) (TWord (S (S (S _)))) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x - y )" := (Op (Sub (TWord (S (S (S _)))) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint8_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - '(uint16_t)' y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - '(uint16_t)' y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - '(uint16_t)' y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - '(uint16_t)' y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint16_t)' x - '(uint16_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair x y)) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint16_t)' y )" := (Op (Sub (TWord 4) (TWord (S (S (S (S _))))) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x - y )" := (Op (Sub (TWord (S (S (S (S _))))) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint16_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - '(uint32_t)' y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - '(uint32_t)' y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - '(uint32_t)' y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - '(uint32_t)' y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint32_t)' x - '(uint32_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair x y)) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint32_t)' y )" := (Op (Sub (TWord 5) (TWord (S (S (S (S (S _)))))) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S _)))))) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint32_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - '(uint64_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - '(uint64_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - '(uint64_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - '(uint64_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint64_t)' x - '(uint64_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair x y)) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint64_t)' y )" := (Op (Sub (TWord 6) (TWord (S (S (S (S (S (S _))))))) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S _))))))) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint64_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - '(uint128_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - '(uint128_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - '(uint128_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - '(uint128_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint128_t)' x - '(uint128_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair x y)) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint128_t)' y )" := (Op (Sub (TWord 7) (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S _)))))))) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint128_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - '(uint256_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - '(uint256_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - '(uint256_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - '(uint256_t)' y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9, format "( '(uint256_t)' x - '(uint256_t)' y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x y)) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair x y)) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair x (Var y))) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) y)) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - '(uint256_t)' y )" := (Op (Sub (TWord 8) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, y at level 9, format "( x - '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x - y )" := (Op (Sub (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, format "( '(uint256_t)' x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (Pair x y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x - y )") : expr_scope.
-Notation "( x - y )" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x - y )") : expr_scope.
-Notation "( '1&(' x * y ')' )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40, format "( '1&(' x * y ')' )").
-Notation "( '1&(' x * y ')' )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, format "( '1&(' x * y ')' )").
-Notation "( '1&(' x * y ')' )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, format "( '1&(' x * y ')' )").
-Notation "( '1&(' x * y ')' )" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, format "( '1&(' x * y ')' )").
-Notation "( '1&(' x + y ')' )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, format "( '1&(' x + y ')' )").
-Notation "( '1&(' x + y ')' )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, format "( '1&(' x + y ')' )").
-Notation "( '1&(' x + y ')' )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, format "( '1&(' x + y ')' )").
-Notation "( '1&(' x + y ')' )" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, format "( '1&(' x + y ')' )").
-Notation "( '1&(' x - y ')' )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 50, format "( '1&(' x - y ')' )").
-Notation "( '1&(' x - y ')' )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, format "( '1&(' x - y ')' )").
-Notation "( '1&(' x - y ')' )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, format "( '1&(' x - y ')' )").
-Notation "( '1&(' x - y ')' )" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, format "( '1&(' x - y ')' )").
-Notation "( x & y )" := (Op (Land _ _ _) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x &ℤ y )" := (Op (Land _ _ TZ) (Pair x y)) (at level 40, format "( x &ℤ y )") : expr_scope.
-Notation "( x & y )" := (Op (Land _ _ _) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x &ℤ y )" := (Op (Land _ _ TZ) (Pair x (Var y))) (at level 40, format "( x &ℤ y )") : expr_scope.
-Notation "( x & y )" := (Op (Land _ _ _) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x &ℤ y )" := (Op (Land _ _ TZ) (Pair (Var x) y)) (at level 40, format "( x &ℤ y )") : expr_scope.
-Notation "( x & y )" := (Op (Land _ _ _) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x &ℤ y )" := (Op (Land _ _ TZ) (Pair (Var x) (Var y))) (at level 40, format "( x &ℤ y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40, y at level 9, format "( x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & y )" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x & y )") : expr_scope.
-Notation "( x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & y )" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x & y )") : expr_scope.
-Notation "( x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & y )" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x & y )") : expr_scope.
-Notation "( x & '1&(uint8_t/*bool*/)' y )" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x & y )" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '1&(uint8_t/*bool*/)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & '(uint8_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint8_t)' x & '(uint8_t)' y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & '(uint8_t)' y )" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x & y )" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint8_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint16_t)' x & '(uint16_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x & '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x & '(uint16_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x & '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x & '(uint16_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x & '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x & '(uint16_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint16_t)' x & '(uint16_t)' y )") : expr_scope.
-Notation "( x & '(uint16_t)' y )" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x & y )" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 40, x at level 9, format "( '(uint16_t)' x & y )") : expr_scope.
-Notation "( x & '(uint16_t)' y )" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x & y )" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint16_t)' x & y )") : expr_scope.
-Notation "( x & '(uint16_t)' y )" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x & y )" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint16_t)' x & y )") : expr_scope.
-Notation "( x & '(uint16_t)' y )" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x & y )" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint16_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint32_t)' x & '(uint32_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x & '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x & '(uint32_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x & '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x & '(uint32_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x & '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x & '(uint32_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint32_t)' x & '(uint32_t)' y )") : expr_scope.
-Notation "( x & '(uint32_t)' y )" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x & y )" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 40, x at level 9, format "( '(uint32_t)' x & y )") : expr_scope.
-Notation "( x & '(uint32_t)' y )" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x & y )" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint32_t)' x & y )") : expr_scope.
-Notation "( x & '(uint32_t)' y )" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x & y )" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint32_t)' x & y )") : expr_scope.
-Notation "( x & '(uint32_t)' y )" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x & y )" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint32_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint64_t)' x & '(uint64_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x & '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x & '(uint64_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x & '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x & '(uint64_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x & '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x & '(uint64_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint64_t)' x & '(uint64_t)' y )") : expr_scope.
-Notation "( x & '(uint64_t)' y )" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x & y )" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 40, x at level 9, format "( '(uint64_t)' x & y )") : expr_scope.
-Notation "( x & '(uint64_t)' y )" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x & y )" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint64_t)' x & y )") : expr_scope.
-Notation "( x & '(uint64_t)' y )" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x & y )" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint64_t)' x & y )") : expr_scope.
-Notation "( x & '(uint64_t)' y )" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x & y )" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint64_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint128_t)' x & '(uint128_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x & '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x & '(uint128_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x & '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x & '(uint128_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x & '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x & '(uint128_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint128_t)' x & '(uint128_t)' y )") : expr_scope.
-Notation "( x & '(uint128_t)' y )" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x & y )" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 40, x at level 9, format "( '(uint128_t)' x & y )") : expr_scope.
-Notation "( x & '(uint128_t)' y )" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x & y )" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint128_t)' x & y )") : expr_scope.
-Notation "( x & '(uint128_t)' y )" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x & y )" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint128_t)' x & y )") : expr_scope.
-Notation "( x & '(uint128_t)' y )" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x & y )" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint128_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( '(uint256_t)' x & '(uint256_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x & '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x & '(uint256_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x & '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x & '(uint256_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x & '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x & '(uint256_t)' y )" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9, format "( '(uint256_t)' x & '(uint256_t)' y )") : expr_scope.
-Notation "( x & '(uint256_t)' y )" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (at level 40, y at level 9, format "( x & '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x & y )" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (at level 40, x at level 9, format "( '(uint256_t)' x & y )") : expr_scope.
-Notation "( x & '(uint256_t)' y )" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 40, y at level 9, format "( x & '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x & y )" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9, format "( '(uint256_t)' x & y )") : expr_scope.
-Notation "( x & '(uint256_t)' y )" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 40, y at level 9, format "( x & '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x & y )" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9, format "( '(uint256_t)' x & y )") : expr_scope.
-Notation "( x & '(uint256_t)' y )" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, y at level 9, format "( x & '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x & y )" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9, format "( '(uint256_t)' x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (Pair x y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x & y )") : expr_scope.
-Notation "( x & y )" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x & y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor _ _ _) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x |ℤ y )" := (Op (Lor _ _ TZ) (Pair x y)) (at level 45, format "( x |ℤ y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor _ _ _) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x |ℤ y )" := (Op (Lor _ _ TZ) (Pair x (Var y))) (at level 45, format "( x |ℤ y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor _ _ _) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x |ℤ y )" := (Op (Lor _ _ TZ) (Pair (Var x) y)) (at level 45, format "( x |ℤ y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor _ _ _) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x |ℤ y )" := (Op (Lor _ _ TZ) (Pair (Var x) (Var y))) (at level 45, format "( x |ℤ y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '1&(uint8_t/*bool*/)' x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 45, y at level 9, format "( x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | y )" := (Op (Lor (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 45, x at level 9, format "( '1&(uint8_t/*bool*/)' x | y )") : expr_scope.
-Notation "( x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | y )" := (Op (Lor (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 45, x at level 9, format "( '1&(uint8_t/*bool*/)' x | y )") : expr_scope.
-Notation "( x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | y )" := (Op (Lor (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '1&(uint8_t/*bool*/)' x | y )") : expr_scope.
-Notation "( x | '1&(uint8_t/*bool*/)' y )" := (Op (Lor (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '1&(uint8_t/*bool*/)' y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x | y )" := (Op (Lor (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '1&(uint8_t/*bool*/)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 0) (TWord 0) (TWord 0)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 0) (TWord 0) (TWord 0)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 0) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 1) (TWord 1) (TWord 1)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 1) (TWord 1) (TWord 1)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 1) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 2) (TWord 2) (TWord 2)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 2) (TWord 2) (TWord 2)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 2) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | '(uint8_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint8_t)' x | '(uint8_t)' y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | '(uint8_t)' y )" := (Op (Lor (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint8_t)' y )") : expr_scope.
-Notation "( '(uint8_t)' x | y )" := (Op (Lor (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint8_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 3) (TWord 3) (TWord 3)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 3) (TWord 3) (TWord 3)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 3) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint16_t)' x | '(uint16_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint16_t)' x | '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x | '(uint16_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint16_t)' x | '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x | '(uint16_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint16_t)' x | '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x | '(uint16_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint16_t)' x | '(uint16_t)' y )") : expr_scope.
-Notation "( x | '(uint16_t)' y )" := (Op (Lor (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x | y )" := (Op (Lor (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 45, x at level 9, format "( '(uint16_t)' x | y )") : expr_scope.
-Notation "( x | '(uint16_t)' y )" := (Op (Lor (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x | y )" := (Op (Lor (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint16_t)' x | y )") : expr_scope.
-Notation "( x | '(uint16_t)' y )" := (Op (Lor (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x | y )" := (Op (Lor (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint16_t)' x | y )") : expr_scope.
-Notation "( x | '(uint16_t)' y )" := (Op (Lor (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint16_t)' y )") : expr_scope.
-Notation "( '(uint16_t)' x | y )" := (Op (Lor (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint16_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 4) (TWord 4) (TWord 4)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 4) (TWord 4) (TWord 4)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 4) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint32_t)' x | '(uint32_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint32_t)' x | '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x | '(uint32_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint32_t)' x | '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x | '(uint32_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint32_t)' x | '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x | '(uint32_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint32_t)' x | '(uint32_t)' y )") : expr_scope.
-Notation "( x | '(uint32_t)' y )" := (Op (Lor (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x | y )" := (Op (Lor (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 45, x at level 9, format "( '(uint32_t)' x | y )") : expr_scope.
-Notation "( x | '(uint32_t)' y )" := (Op (Lor (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x | y )" := (Op (Lor (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint32_t)' x | y )") : expr_scope.
-Notation "( x | '(uint32_t)' y )" := (Op (Lor (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x | y )" := (Op (Lor (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint32_t)' x | y )") : expr_scope.
-Notation "( x | '(uint32_t)' y )" := (Op (Lor (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint32_t)' y )") : expr_scope.
-Notation "( '(uint32_t)' x | y )" := (Op (Lor (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint32_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 5) (TWord 5) (TWord 5)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 5) (TWord 5) (TWord 5)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 5) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint64_t)' x | '(uint64_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint64_t)' x | '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x | '(uint64_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint64_t)' x | '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x | '(uint64_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint64_t)' x | '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x | '(uint64_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint64_t)' x | '(uint64_t)' y )") : expr_scope.
-Notation "( x | '(uint64_t)' y )" := (Op (Lor (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x | y )" := (Op (Lor (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 45, x at level 9, format "( '(uint64_t)' x | y )") : expr_scope.
-Notation "( x | '(uint64_t)' y )" := (Op (Lor (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x | y )" := (Op (Lor (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint64_t)' x | y )") : expr_scope.
-Notation "( x | '(uint64_t)' y )" := (Op (Lor (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x | y )" := (Op (Lor (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint64_t)' x | y )") : expr_scope.
-Notation "( x | '(uint64_t)' y )" := (Op (Lor (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint64_t)' y )") : expr_scope.
-Notation "( '(uint64_t)' x | y )" := (Op (Lor (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint64_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 6) (TWord 6) (TWord 6)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 6) (TWord 6) (TWord 6)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 6) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint128_t)' x | '(uint128_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint128_t)' x | '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x | '(uint128_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint128_t)' x | '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x | '(uint128_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint128_t)' x | '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x | '(uint128_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint128_t)' x | '(uint128_t)' y )") : expr_scope.
-Notation "( x | '(uint128_t)' y )" := (Op (Lor (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x | y )" := (Op (Lor (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 45, x at level 9, format "( '(uint128_t)' x | y )") : expr_scope.
-Notation "( x | '(uint128_t)' y )" := (Op (Lor (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x | y )" := (Op (Lor (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint128_t)' x | y )") : expr_scope.
-Notation "( x | '(uint128_t)' y )" := (Op (Lor (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x | y )" := (Op (Lor (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint128_t)' x | y )") : expr_scope.
-Notation "( x | '(uint128_t)' y )" := (Op (Lor (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint128_t)' y )") : expr_scope.
-Notation "( '(uint128_t)' x | y )" := (Op (Lor (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint128_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 7) (TWord 7) (TWord 7)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 7) (TWord 7) (TWord 7)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 7) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( '(uint256_t)' x | '(uint256_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 45, x at level 9, y at level 9, format "( '(uint256_t)' x | '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x | '(uint256_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint256_t)' x | '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x | '(uint256_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 45, x at level 9, y at level 9, format "( '(uint256_t)' x | '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x | '(uint256_t)' y )" := (Op (Lor (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 45, x at level 9, y at level 9, format "( '(uint256_t)' x | '(uint256_t)' y )") : expr_scope.
-Notation "( x | '(uint256_t)' y )" := (Op (Lor (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (at level 45, y at level 9, format "( x | '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x | y )" := (Op (Lor (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (at level 45, x at level 9, format "( '(uint256_t)' x | y )") : expr_scope.
-Notation "( x | '(uint256_t)' y )" := (Op (Lor (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 45, y at level 9, format "( x | '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x | y )" := (Op (Lor (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 45, x at level 9, format "( '(uint256_t)' x | y )") : expr_scope.
-Notation "( x | '(uint256_t)' y )" := (Op (Lor (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 45, y at level 9, format "( x | '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x | y )" := (Op (Lor (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 45, x at level 9, format "( '(uint256_t)' x | y )") : expr_scope.
-Notation "( x | '(uint256_t)' y )" := (Op (Lor (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 45, y at level 9, format "( x | '(uint256_t)' y )") : expr_scope.
-Notation "( '(uint256_t)' x | y )" := (Op (Lor (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 45, x at level 9, format "( '(uint256_t)' x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 8) (TWord 8) (TWord 8)) (Pair x y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)) (format "( x | y )") : expr_scope.
-Notation "( x | y )" := (Op (Lor (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (format "( x | y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl _ _ _) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( x <<ℤ y )" := (Op (Shl _ _ TZ) (Pair x y)) (at level 30, format "( x <<ℤ y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl _ _ _) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( x <<ℤ y )" := (Op (Shl _ _ TZ) (Pair x (Var y))) (at level 30, format "( x <<ℤ y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl _ _ _) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( x <<ℤ y )" := (Op (Shl _ _ TZ) (Pair (Var x) y)) (at level 30, format "( x <<ℤ y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl _ _ _) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( x <<ℤ y )" := (Op (Shl _ _ TZ) (Pair (Var x) (Var y))) (at level 30, format "( x <<ℤ y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 30, format "( '1&(uint8_t/*bool*/)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30, format "( '1&(uint8_t/*bool*/)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30, format "( '1&(uint8_t/*bool*/)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '1&(uint8_t/*bool*/)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30, format "( '1&(uint8_t/*bool*/)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint8_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint8_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint16_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 30, format "( '(uint16_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint16_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30, format "( '(uint16_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint16_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30, format "( '(uint16_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint16_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint16_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint32_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 30, format "( '(uint32_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint32_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30, format "( '(uint32_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint32_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30, format "( '(uint32_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint32_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint32_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint64_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 30, format "( '(uint64_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint64_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30, format "( '(uint64_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint64_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30, format "( '(uint64_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint64_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint64_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint128_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 30, format "( '(uint128_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint128_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30, format "( '(uint128_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint128_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30, format "( '(uint128_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint128_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint128_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint256_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 30, format "( '(uint256_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint256_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 30, format "( '(uint256_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (format "( x << y )") : expr_scope.
-Notation "( '(uint256_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 30, format "( '(uint256_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (format "( x << y )") : expr_scope.
-Notation "( '(uint256_t)' x << y )" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 30, format "( '(uint256_t)' x << y )") : expr_scope.
-Notation "( x << y )" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (format "( x << y )") : expr_scope.
-Notation "( x >> y )" := (Op (Shr _ _ _) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "( x >>ℤ y )" := (Op (Shr _ _ TZ) (Pair x y)) (at level 30, format "( x >>ℤ y )") : expr_scope.
-Notation "( x >> y )" := (Op (Shr _ _ _) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "( x >>ℤ y )" := (Op (Shr _ _ TZ) (Pair x (Var y))) (at level 30, format "( x >>ℤ y )") : expr_scope.
-Notation "( x >> y )" := (Op (Shr _ _ _) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "( x >>ℤ y )" := (Op (Shr _ _ TZ) (Pair (Var x) y)) (at level 30, format "( x >>ℤ y )") : expr_scope.
-Notation "( x >> y )" := (Op (Shr _ _ _) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "( x >>ℤ y )" := (Op (Shr _ _ TZ) (Pair (Var x) (Var y))) (at level 30, format "( x >>ℤ y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint8_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint16_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint32_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint64_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint128_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (format "( x >> y )") : expr_scope.
-Notation "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 30) : expr_scope.
-Notation "( x >> y )" := (Op (Shr (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (format "( x >> y )") : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz32' ( v , x , y )") : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S _)))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S _)))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz32' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S _)))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz64' ( v , x , y )") : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz64' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S _))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznz128' ( v , x , y )") : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznz128' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _)))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair v x) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair v x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair v (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair (Var v) x) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _)) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t/*bool*/)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S _))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S _)))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair v x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair v x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint8_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S _))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint8_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v x) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint16_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S _)))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint16_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v x) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint32_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S _))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint32_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v x) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint64_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S _)))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint64_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v x) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair v (Var x)) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) x) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) y)) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'(uint128_t)' 'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S _))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) (format "'(uint128_t)' 'cmovznzℤ' ( v , x , y )") : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord _) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect (TWord _) (TWord _) (TWord (S (S (S (S (S (S (S (S _))))))))) (TWord (S (S (S (S (S (S (S (S _)))))))))) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair v x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair v x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair v (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair v (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair (Var v) x) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair (Var v) x) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair (Var v) (Var x)) y)) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect TZ _ _ _) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ TZ _ _) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ TZ _) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'cmovznzℤ' ( v , x , y )" := (Op (Zselect _ _ _ TZ) (Pair (Pair (Var v) (Var x)) (Var y))) : expr_scope.
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32' ( c , a , b )" := (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64' ( c , a , b )" := (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128' ( c , a , b )" := (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25' ( c , a , b )" := (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26' ( c , a , b )" := (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51' ( c , a , b )" := (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32' ( c , a , b )" := (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64' ( c , a , b )" := (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 1) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 2) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 3) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 4) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 5) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 6) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 1) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 2) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 3) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 4) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 5) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 6) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128' ( c , a , b )" := (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25' ( c , a , b )" := (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 1) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 2) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 3) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 4) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 1) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 2) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 3) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 4) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26' ( c , a , b )" := (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 (TWord 0) (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 1) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 2) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 3) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 4) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 5) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 1) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 2) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 3) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 4) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 5) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51' ( c , a , b )" := (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u1' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 0) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 0) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 0) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32_out_u8' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 3) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 3) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 3) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u1' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64_out_u8' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u1' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 0) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 0) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 0) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128_out_u8' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 3) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 3) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 3) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'1&(uint8_t/*bool*/)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u1' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 0) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 0) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '1&(uint8_t/*bool*/)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 0) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u1' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 0)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'(uint8_t)' 'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51_out_u8' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 3) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 3) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '(uint8_t)' '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 3) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51_out_u8' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 3)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32' ( a , b )" := (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 (TWord 5) (TWord 5) (TWord 5) (TWord 5)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64' ( a , b )" := (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128' ( a , b )" := (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 (TWord 7) (TWord 7) (TWord 7) (TWord 7)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51' ( a , b )" := (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 (TWord 6) (TWord 6) (TWord 6) (TWord 6)) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u32ℤ' ( c , a , b )" := (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u64ℤ' ( c , a , b )" := (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u128ℤ' ( c , a , b )" := (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u25ℤ' ( c , a , b )" := (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u26ℤ' ( c , a , b )" := (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'addcarryx_u51ℤ' ( c , a , b )" := (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_addcarryx_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (AddWithGetCarry 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u32ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u32ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 32 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u64ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u64ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 64 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u128ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u128ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 128 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u25ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u25ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 25 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u26ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u26ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 26 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c a) b)) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair c (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-Notation "'subborrow_u51ℤ' ( c , a , b )" := (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ _ TZ) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_subborrow_u51ℤ' ( c , a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (SubWithGetBorrow 51 _ _ _ TZ _) (Pair (Pair (Var c) (Var a)) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ _ TZ) (Pair a b)) : expr_scope.
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ TZ _) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ _ TZ) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ TZ _) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ _ TZ) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ TZ _) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ _ TZ) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ TZ _) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ _ TZ) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ TZ _) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ _ TZ) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ TZ _) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ _ TZ) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u32ℤ' ( a , b )" := (Op (MulSplit 32 _ _ TZ _) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ _ TZ) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u32ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 32 _ _ TZ _) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ _ TZ) (Pair a b)) : expr_scope.
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ TZ _) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ _ TZ) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ TZ _) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ _ TZ) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ TZ _) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ _ TZ) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ TZ _) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ _ TZ) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ TZ _) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ _ TZ) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ TZ _) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ _ TZ) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u64ℤ' ( a , b )" := (Op (MulSplit 64 _ _ TZ _) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ _ TZ) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u64ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 64 _ _ TZ _) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ _ TZ) (Pair a b)) : expr_scope.
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ TZ _) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ _ TZ) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ TZ _) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ _ TZ) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ TZ _) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ _ TZ) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ TZ _) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ _ TZ) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ TZ _) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ _ TZ) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ TZ _) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ _ TZ) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u128ℤ' ( a , b )" := (Op (MulSplit 128 _ _ TZ _) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ _ TZ) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u128ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 128 _ _ TZ _) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ _ TZ) (Pair a b)) : expr_scope.
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ TZ _) (Pair a b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ _ TZ) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ TZ _) (Pair a b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ _ TZ) (Pair a (Var b))) : expr_scope.
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ TZ _) (Pair a (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ _ TZ) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ TZ _) (Pair a (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ _ TZ) (Pair (Var a) b)) : expr_scope.
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ TZ _) (Pair (Var a) b)) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ _ TZ) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ TZ _) (Pair (Var a) b)) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ _ TZ) (Pair (Var a) (Var b))) : expr_scope.
-Notation "'mulx_u51ℤ' ( a , b )" := (Op (MulSplit 51 _ _ TZ _) (Pair (Var a) (Var b))) : expr_scope.
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ _ TZ) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-(*Notation "T0 out ; T1 c_out = '_mulx_u51ℤ' ( a , b , & out ) ; REST" := (LetIn (tx:=Prod T0 T1) (Op (MulSplit 51 _ _ TZ _) (Pair (Var a) (Var b))) (fun '((out, c_out)%core) => REST)) : expr_scope.*)
-Notation Return x := (Var x).
-Notation C_like := (Expr base_type op _).
diff --git a/src/Compilers/Z/CommonSubexpressionElimination.v b/src/Compilers/Z/CommonSubexpressionElimination.v
deleted file mode 100644
index 2493f6041..000000000
--- a/src/Compilers/Z/CommonSubexpressionElimination.v
+++ /dev/null
@@ -1,230 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.Sorting.Mergesort.
-Require Import Coq.Structures.Orders.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Util.
-Require Import Crypto.Compilers.CommonSubexpressionElimination.
-Require Import Crypto.Compilers.CommonSubexpressionEliminationProperties.
-
-Local Set Decidable Equality Schemes.
-Local Set Boolean Equality Schemes.
-Inductive symbolic_op :=
-| SOpConst (z : Z)
-| SAdd
-| SSub
-| SMul
-| SShl
-| SShr
-| SLand
-| SLor
-| SOpp
-| SIdWithAlt
-| SZselect
-| SMulSplit (bitwidth : Z)
-| SAddWithCarry
-| SAddWithGetCarry (bitwidth : Z)
-| SSubWithBorrow
-| SSubWithGetBorrow (bitwidth : Z)
-.
-
-Definition symbolic_op_leb (x y : symbolic_op) : bool
- := match x, y with
- | SOpConst z1, SOpConst z2 => Z.leb z1 z2
- | SAddWithGetCarry bw1, SAddWithGetCarry bw2 => Z.leb bw1 bw2
- | SSubWithGetBorrow bw1, SSubWithGetBorrow bw2 => Z.leb bw1 bw2
- | SOpConst _, _ => true
- | _, SOpConst _ => false
- | SAdd, _ => true
- | _, SAdd => false
- | SSub, _ => true
- | _, SSub => false
- | SMul, _ => true
- | _, SMul => false
- | SShl, _ => true
- | _, SShl => false
- | SShr, _ => true
- | _, SShr => false
- | SLand, _ => true
- | _, SLand => false
- | SLor, _ => true
- | _, SLor => false
- | SOpp, _ => true
- | _, SOpp => false
- | SIdWithAlt, _ => true
- | _, SIdWithAlt => false
- | SZselect, _ => true
- | _, SZselect => false
- | SMulSplit _, _ => true
- | _, SMulSplit _ => false
- | SAddWithCarry, _ => true
- | _, SAddWithCarry => false
- | SAddWithGetCarry _, _ => true
- | _, SAddWithGetCarry _ => false
- | SSubWithBorrow, _ => true
- | _, SSubWithBorrow => false
- (*| SSubWithGetBorrow _, _ => true
- | _, SSubWithGetBorrow _ => false*)
- end.
-
-Local Notation symbolic_expr := (@symbolic_expr base_type symbolic_op).
-Local Notation symbolic_expr_beq := (@symbolic_expr_beq base_type symbolic_op base_type_beq symbolic_op_beq).
-Local Notation symbolic_expr_leb := (@symbolic_expr_leb base_type symbolic_op base_type_beq symbolic_op_beq symbolic_op_leb base_type_leb).
-
-Definition symbolize_op s d (opc : op s d) : symbolic_op
- := match opc with
- | OpConst T z => SOpConst z
- | Add T1 T2 Tout => SAdd
- | Sub T1 T2 Tout => SSub
- | Mul T1 T2 Tout => SMul
- | Shl T1 T2 Tout => SShl
- | Shr T1 T2 Tout => SShr
- | Land T1 T2 Tout => SLand
- | Lor T1 T2 Tout => SLor
- | Opp T Tout => SOpp
- | IdWithAlt T1 T2 Tout => SIdWithAlt
- | Zselect T1 T2 T3 Tout => SZselect
- | MulSplit bitwidth T1 T2 Tout1 Tout2 => SMulSplit bitwidth
- | AddWithCarry T1 T2 T3 Tout => SAddWithCarry
- | AddWithGetCarry bitwidth T1 T2 T3 Tout1 Tout2 => SAddWithGetCarry bitwidth
- | SubWithBorrow T1 T2 T3 Tout => SSubWithBorrow
- | SubWithGetBorrow bitwidth T1 T2 T3 Tout1 Tout2 => SSubWithGetBorrow bitwidth
- end.
-
-Definition denote_symbolic_op s d (opc : symbolic_op) : option (op s d)
- := match opc, s, d with
- | SOpConst z, Unit, Tbase T => Some (OpConst z)
- | SAdd, Prod (Tbase _) (Tbase _), Tbase _ => Some (Add _ _ _)
- | SSub, Prod (Tbase _) (Tbase _), Tbase _ => Some (Sub _ _ _)
- | SMul, Prod (Tbase _) (Tbase _), Tbase _ => Some (Mul _ _ _)
- | SShl, Prod (Tbase _) (Tbase _), Tbase _ => Some (Shl _ _ _)
- | SShr, Prod (Tbase _) (Tbase _), Tbase _ => Some (Shr _ _ _)
- | SLand, Prod (Tbase _) (Tbase _), Tbase _ => Some (Land _ _ _)
- | SLor, Prod (Tbase _) (Tbase _), Tbase _ => Some (Lor _ _ _)
- | SOpp, Tbase _, Tbase _ => Some (Opp _ _)
- | SIdWithAlt, Prod (Tbase _) (Tbase _), Tbase _ => Some (IdWithAlt _ _ _)
- | SZselect, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Tbase _ => Some (Zselect _ _ _ _)
- | SMulSplit bitwidth, Prod (Tbase _) (Tbase _), Prod (Tbase _) (Tbase _)
- => Some (MulSplit bitwidth _ _ _ _)
- | SAddWithCarry, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Tbase _ => Some (AddWithCarry _ _ _ _)
- | SAddWithGetCarry bitwidth, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Prod (Tbase _) (Tbase _)
- => Some (AddWithGetCarry bitwidth _ _ _ _ _)
- | SSubWithBorrow, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Tbase _ => Some (SubWithBorrow _ _ _ _)
- | SSubWithGetBorrow bitwidth, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Prod (Tbase _) (Tbase _)
- => Some (SubWithGetBorrow bitwidth _ _ _ _ _)
- | SAdd, _, _
- | SSub, _, _
- | SMul, _, _
- | SShl, _, _
- | SShr, _, _
- | SLand, _, _
- | SLor, _, _
- | SOpp, _, _
- | SOpConst _, _, _
- | SIdWithAlt, _, _
- | SZselect, _, _
- | SMulSplit _, _, _
- | SAddWithCarry, _, _
- | SAddWithGetCarry _, _, _
- | SSubWithBorrow, _, _
- | SSubWithGetBorrow _, _, _
- => None
- end.
-
-Lemma symbolic_op_leb_total
- : forall a1 a2, symbolic_op_leb a1 a2 = true \/ symbolic_op_leb a2 a1 = true.
-Proof.
- induction a1, a2; simpl; auto;
- rewrite !Z.leb_le; omega.
-Qed.
-
-Module SymbolicExprOrder <: TotalLeBool.
- Definition t := (flat_type base_type * symbolic_expr)%type.
- Definition leb (x y : t) : bool := symbolic_expr_leb (snd x) (snd y).
- Theorem leb_total : forall a1 a2, leb a1 a2 = true \/ leb a2 a1 = true.
- Proof.
- intros; apply symbolic_expr_leb_total;
- auto using internal_base_type_dec_bl, internal_base_type_dec_lb, internal_symbolic_op_dec_bl, internal_symbolic_op_dec_lb, base_type_leb_total, symbolic_op_leb_total.
- Qed.
-End SymbolicExprOrder.
-
-Module Import SymbolicExprSort := Sort SymbolicExprOrder.
-
-Fixpoint symbolic_op_args_to_list (t : flat_type base_type)
- (opc : symbolic_op) (args : symbolic_expr)
- : list (flat_type base_type * symbolic_expr)
- := match args, t with
- | SOp argT opc' args', _
- => if symbolic_op_beq opc opc'
- then symbolic_op_args_to_list argT opc args'
- else (t, args)::nil
- | SPair x y, Prod A B
- => symbolic_op_args_to_list A opc x ++ symbolic_op_args_to_list B opc y
- | SPair x y, Unit
- => symbolic_op_args_to_list Unit opc x ++ symbolic_op_args_to_list Unit opc y
- | STT, _
- | SVar _, _
- | SPair _ _, _
- | SFst _ _ _, _
- | SSnd _ _ _, _
- | SInvalid, _
- => (t, args)::nil
- end%list.
-
-Fixpoint symbolic_op_list_to_args (args : list (flat_type base_type * symbolic_expr)) : symbolic_expr
- := match args with
- | nil => SInvalid
- | (t, arg)::nil => arg
- | (t1, arg1)::(t2, arg2)::nil
- => SPair arg1 arg2
- | (t1, arg1)::(((t2, arg2)::args'') as args')
- => SPair arg1 (SOp t2 SAdd (symbolic_op_list_to_args args'))
- end%list.
-
-Definition normalize_symbolic_expr_mod_c (opc : symbolic_op) (args : symbolic_expr) : symbolic_expr
- := match opc with
- | SAdd
- | SMul
- | SLand
- | SLor
- => let ls := symbolic_op_args_to_list Unit opc args in
- let ls := sort ls in
- symbolic_op_list_to_args ls
- | SOpConst _
- | SSub
- | SShl
- | SShr
- | SOpp
- | SIdWithAlt
- | SZselect
- | SMulSplit _
- | SAddWithCarry
- | SAddWithGetCarry _
- | SSubWithBorrow
- | SSubWithGetBorrow _
- => args
- end.
-
-Definition csef inline_symbolic_expr_in_lookup {var t} (v : exprf _ _ t) xs
- := @csef base_type symbolic_op base_type_beq symbolic_op_beq
- internal_base_type_dec_bl op symbolize_op
- normalize_symbolic_expr_mod_c
- var inline_symbolic_expr_in_lookup t v xs.
-
-Definition cse inline_symbolic_expr_in_lookup {var} (prefix : list _) {t} (v : expr _ _ t) xs
- := @cse base_type symbolic_op base_type_beq symbolic_op_beq
- internal_base_type_dec_bl op symbolize_op
- normalize_symbolic_expr_mod_c
- inline_symbolic_expr_in_lookup var prefix t v xs.
-
-Definition CSE_gen inline_symbolic_expr_in_lookup {t} (e : Expr t) (prefix : forall var, list { t : flat_type base_type & exprf _ _ t })
- : Expr t
- := @CSE base_type symbolic_op base_type_beq symbolic_op_beq
- internal_base_type_dec_bl op symbolize_op
- normalize_symbolic_expr_mod_c
- inline_symbolic_expr_in_lookup t e prefix.
-
-Definition CSE inline_symbolic_expr_in_lookup {t} (e : Expr t)
- : Expr t
- := @CSE_gen inline_symbolic_expr_in_lookup t e (fun _ => nil).
diff --git a/src/Compilers/Z/CommonSubexpressionEliminationInterp.v b/src/Compilers/Z/CommonSubexpressionEliminationInterp.v
deleted file mode 100644
index 7e7f7910a..000000000
--- a/src/Compilers/Z/CommonSubexpressionEliminationInterp.v
+++ /dev/null
@@ -1,23 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Coq.Lists.List.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.CommonSubexpressionEliminationInterp.
-Require Import Crypto.Compilers.Z.CommonSubexpressionElimination.
-
-Lemma InterpCSE_gen inline_symbolic_expr_in_lookup t (e : Expr t) prefix
- (Hwf : Wf e)
- : forall x, Interp (@CSE_gen inline_symbolic_expr_in_lookup t e prefix) x = Interp e x.
-Proof.
- apply InterpCSE;
- auto using internal_base_type_dec_bl, internal_base_type_dec_lb, internal_symbolic_op_dec_bl, internal_symbolic_op_dec_lb, denote_symbolic_op.
-Qed.
-
-Lemma InterpCSE inline_symbolic_expr_in_lookup t (e : Expr t) (Hwf : Wf e)
- : forall x, Interp (@CSE inline_symbolic_expr_in_lookup t e) x = Interp e x.
-Proof.
- apply InterpCSE_gen; auto.
-Qed.
-
-Hint Rewrite @InterpCSE using solve_wf_side_condition : reflective_interp.
diff --git a/src/Compilers/Z/CommonSubexpressionEliminationWf.v b/src/Compilers/Z/CommonSubexpressionEliminationWf.v
deleted file mode 100644
index 93d422b2d..000000000
--- a/src/Compilers/Z/CommonSubexpressionEliminationWf.v
+++ /dev/null
@@ -1,29 +0,0 @@
-(** * Common Subexpression Elimination for PHOAS Syntax *)
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.CommonSubexpressionEliminationWf.
-Require Import Crypto.Compilers.Z.CommonSubexpressionElimination.
-
-Lemma Wf_CSE_gen inline_symbolic_expr_in_lookup t (e : Expr t)
- prefix
- (Hlen : forall var1 var2, length (prefix var1) = length (prefix var2))
- (Hprefix : forall var1 var2 n t1 t2 e1 e2,
- List.nth_error (prefix var1) n = Some (existT _ t1 e1)
- -> List.nth_error (prefix var2) n = Some (existT _ t2 e2)
- -> exists pf : t1 = t2, wff nil (eq_rect _ (@exprf _ _ _) e1 _ pf) e2)
- (Hwf : Wf e)
- : Wf (@CSE_gen inline_symbolic_expr_in_lookup t e prefix).
-Proof.
- apply Wf_CSE; auto using internal_base_type_dec_bl, internal_base_type_dec_lb, internal_symbolic_op_dec_bl, internal_symbolic_op_dec_lb.
-Qed.
-
-Lemma Wf_CSE inline_symbolic_expr_in_lookup t (e : Expr t)
- (Hwf : Wf e)
- : Wf (@CSE inline_symbolic_expr_in_lookup t e).
-Proof.
- apply Wf_CSE_gen; simpl; auto.
- { destruct n; simpl; try congruence. }
-Qed.
-
-Hint Resolve Wf_CSE : wf.
diff --git a/src/Compilers/Z/FoldTypes.v b/src/Compilers/Z/FoldTypes.v
deleted file mode 100644
index 7b3059e43..000000000
--- a/src/Compilers/Z/FoldTypes.v
+++ /dev/null
@@ -1,17 +0,0 @@
-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 t) : base_type
- := TypeFold (fun t => t) f init e.
-End min_or_max.
-
-Definition MaxTypeUsed {t} (e : Expr t) : base_type
- := TypeFold base_type_max (TWord 0) e.
-Definition MinTypeUsed {t} (e : Expr t) : base_type
- := TypeFold base_type_min TZ e.
diff --git a/src/Compilers/Z/GeneralizeVar.v b/src/Compilers/Z/GeneralizeVar.v
deleted file mode 100644
index 5431dd77a..000000000
--- a/src/Compilers/Z/GeneralizeVar.v
+++ /dev/null
@@ -1,15 +0,0 @@
-(** * Generalize [var] in [exprf] *)
-Require Import Coq.ZArith.BinInt.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.GeneralizeVar.
-Require Import Crypto.Compilers.Z.Syntax.
-
-(** N.B. This procedure only works when there are no nested lets,
- i.e., nothing like [let x := let y := z in w] in the PHOAS syntax
- tree. This is a limitation of [compile]. *)
-
-Definition GeneralizeVar {t} (e : @Syntax.expr base_type op _ t)
- : option (@Z.Syntax.Expr (domain t -> codomain t))
- := @GeneralizeVar base_type op base_type_beq internal_base_type_dec_bl
- (fun _ t => Op (OpConst 0%Z) TT)
- t e.
diff --git a/src/Compilers/Z/GeneralizeVarInterp.v b/src/Compilers/Z/GeneralizeVarInterp.v
deleted file mode 100644
index a361e24fd..000000000
--- a/src/Compilers/Z/GeneralizeVarInterp.v
+++ /dev/null
@@ -1,23 +0,0 @@
-Require Import Coq.ZArith.BinInt.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.GeneralizeVarInterp.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.GeneralizeVar.
-
-Definition InterpGeneralizeVar
- {interp_base_type}
- {interp_op}
- {t} (e : Expr t)
- (Hwf : Wf e)
- e'
- (He' : GeneralizeVar (e _) = Some e')
- : forall v, Compilers.Syntax.Interp (interp_base_type:=interp_base_type) interp_op e' v
- = Compilers.Syntax.Interp interp_op e v
- := @InterpGeneralizeVar
- base_type op base_type_beq internal_base_type_dec_bl
- internal_base_type_dec_lb
- (fun _ t => Op (OpConst 0%Z) TT)
- interp_base_type
- interp_op
- t e Hwf e' He'.
diff --git a/src/Compilers/Z/GeneralizeVarWf.v b/src/Compilers/Z/GeneralizeVarWf.v
deleted file mode 100644
index d0d97e3c9..000000000
--- a/src/Compilers/Z/GeneralizeVarWf.v
+++ /dev/null
@@ -1,30 +0,0 @@
-Require Import Coq.ZArith.BinInt.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.GeneralizeVarWf.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.GeneralizeVar.
-
-Definition Wf_GeneralizeVar
- {t} (e1 : expr base_type op t)
- e'
- (He' : GeneralizeVar e1 = Some e')
- : Wf e'
- := @Wf_GeneralizeVar
- base_type op base_type_beq internal_base_type_dec_bl
- internal_base_type_dec_lb
- (fun _ t => Op (OpConst 0%Z) TT)
- t e1 e' He'.
-
-Definition Wf_GeneralizeVar_arrow
- {s d} (e : expr base_type op (Arrow s d))
- e'
- (He' : GeneralizeVar e = Some e')
- : Wf e'
- := @Wf_GeneralizeVar_arrow
- base_type op base_type_beq internal_base_type_dec_bl
- internal_base_type_dec_lb
- (fun _ t => Op (OpConst 0%Z) TT)
- s d e e' He'.
-
-Hint Resolve Wf_GeneralizeVar Wf_GeneralizeVar_arrow : wf.
diff --git a/src/Compilers/Z/HexNotationConstants.v b/src/Compilers/Z/HexNotationConstants.v
deleted file mode 100644
index 0e365d9f8..000000000
--- a/src/Compilers/Z/HexNotationConstants.v
+++ /dev/null
@@ -1,7043 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Export Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Export bbv.WordScope.
-Require Export Crypto.Util.Notations.
-
-Notation Const x := (Op (OpConst x) TT).
-(* python:
-<<
-#!/usr/bin/env python
-from __future__ import with_statement
-import math
-
-systematic_nums \
- = (list(range(128)) +
- [2**i + j for i in range(256) for j in (1, 0, -1)])
-nums = tuple(sorted(set(systematic_nums + [
- 187,
- 189,
- 317,
- 421,
- 481,
- 569,
- 765,
- 977,
- 5311,
- 32765,
- 32766,
- 65530,
- 65531,
- 65534,
- 114687,
- 121665,
- 130307,
- 131039,
- 131062,
- 131067,
- 261167,
- 261575,
- 262131,
- 262139,
- 262142,
- 523807,
- 524101,
- 524223,
- 524262,
- 524263,
- 524267,
- 524269,
- 524271,
- 524278,
- 524279,
- 524286,
- 679935,
- 1032192,
- 1044479,
- 1047599,
- 1047811,
- 1048202,
- 1048446,
- 1048471,
- 1048511,
- 1048526,
- 1048534,
- 1048538,
- 1048549,
- 1048557,
- 1048558,
- 1048559,
- 1048573,
- 1048574,
- 1359870,
- 2060031,
- 2064384,
- 2081439,
- 2088958,
- 2094335,
- 2095198,
- 2095622,
- 2096127,
- 2096128,
- 2096583,
- 2096942,
- 2097098,
- 2097114,
- 2097118,
- 2097127,
- 2097135,
- 2097143,
- 2097150,
- 4120062,
- 4162878,
- 4188670,
- 4192254,
- 4193166,
- 4193327,
- 4193539,
- 4193735,
- 4193823,
- 4193883,
- 4193987,
- 4194115,
- 4194240,
- 4194254,
- 4194270,
- 4194271,
- 4194275,
- 4194279,
- 4194285,
- 4194286,
- 4194287,
- 4194289,
- 4194293,
- 4194299,
- 4194301,
- 4194302,
- 8323583,
- 8372224,
- 8386654,
- 8387078,
- 8387470,
- 8387646,
- 8387766,
- 8387974,
- 8388127,
- 8388187,
- 8388230,
- 8388291,
- 8388421,
- 8388491,
- 8388503,
- 8388542,
- 8388558,
- 8388570,
- 8388574,
- 8388577,
- 8388578,
- 8388581,
- 8388591,
- 8388593,
- 8388595,
- 8388598,
- 8388599,
- 8388602,
- 8388603,
- 8388604,
- 8388605,
- 8388606,
- 11599871,
- 16647166,
- 16711679,
- 16775935,
- 16776254,
- 16776374,
- 16776582,
- 16776842,
- 16776959,
- 16776982,
- 16777006,
- 16777027,
- 16777029,
- 16777111,
- 16777162,
- 16777182,
- 16777185,
- 16777186,
- 16777187,
- 16777189,
- 16777190,
- 16777191,
- 16777197,
- 16777198,
- 16777199,
- 16777201,
- 16777205,
- 16777206,
- 16777207,
- 16777210,
- 16777211,
- 16777212,
- 16777213,
- 16777214,
- 23199742,
- 33423358,
- 33551870,
- 33554054,
- 33554058,
- 33554222,
- 33554315,
- 33554370,
- 33554374,
- 33554378,
- 33554382,
- 33554394,
- 33554398,
- 33554399,
- 33554401,
- 33554402,
- 33554407,
- 33554410,
- 33554411,
- 33554413,
- 33554414,
- 33554415,
- 33554417,
- 33554422,
- 33554423,
- 33554426,
- 33554427,
- 33554428,
- 33554429,
- 33554430,
- 67107887,
- 67108630,
- 67108798,
- 67108799,
- 67108802,
- 67108814,
- 67108822,
- 67108826,
- 67108830,
- 67108833,
- 67108834,
- 67108837,
- 67108839,
- 67108843,
- 67108845,
- 67108846,
- 67108847,
- 67108849,
- 67108854,
- 67108855,
- 67108858,
- 67108859,
- 67108860,
- 67108861,
- 67108862,
- 134215774,
- 134217598,
- 134217666,
- 134217674,
- 134217678,
- 134217686,
- 134217690,
- 134217694,
- 134217697,
- 134217698,
- 134217699,
- 134217703,
- 134217709,
- 134217710,
- 134217711,
- 134217713,
- 134217718,
- 134217719,
- 134217722,
- 134217724,
- 134217725,
- 134217726,
- 268431360,
- 268435394,
- 268435398,
- 268435406,
- 268435418,
- 268435422,
- 268435426,
- 268435438,
- 268435441,
- 268435443,
- 268435445,
- 268435447,
- 268435450,
- 268435451,
- 268435452,
- 268435453,
- 268435454,
- 536869935,
- 536870882,
- 536870886,
- 536870890,
- 536870893,
- 536870894,
- 536870902,
- 536870906,
- 536870907,
- 536870908,
- 536870909,
- 536870910,
- 678152731,
- 954437177,
- 1054736383,
- 1065418751,
- 1073709055,
- 1073739870,
- 1073741786,
- 1073741814,
- 1073741818,
- 1073741821,
- 1073741822,
- 1332920885,
- 1749801491,
- 2147418110,
- 2147483631,
- 2147483642,
- 2147483644,
- 2147483646,
- 2147483647,
- 2863311531,
- 2969567231,
- 3123612579,
- 3264175145,
- 3303820997,
- 3435973837,
- 3707621341,
- 4008636143,
- 4042322161,
- 4262789119,
- 4289200127,
- 4294639615,
- 4294901759,
- 4294963199,
- 4294966319,
- 4294966531,
- 4294966727,
- 4294966815,
- 4294966875,
- 4294966979,
- 4294967107,
- 4294967109,
- 4294967179,
- 4294967191,
- 4294967262,
- 4294967263,
- 4294967265,
- 4294967267,
- 4294967269,
- 4294967271,
- 4294967275,
- 4294967277,
- 4294967279,
- 4294967281,
- 4294967283,
- 4294967285,
- 4294967287,
- 4294967291,
- 4294967293,
- 4294967294,
- 4294968273,
- 8589934559,
- 8589934567,
- 8589934575,
- 8589934587,
- 8589934590,
- 17179869134,
- 17179869174,
- 17179869182,
- 34359738341,
- 34359738349,
- 34359738355,
- 34359738366,
- 68719476682,
- 68719476707,
- 68719476710,
- 68719476727,
- 68719476733,
- 68719476734,
- 133143985199,
- 137436856320,
- 137438952991,
- 137438953285,
- 137438953355,
- 137438953454,
- 137438953470,
- 266287970398,
- 274844352511,
- 274877902847,
- 274877906919,
- 274877906927,
- 274877906933,
- 274877906939,
- 274877906941,
- 274877906942,
- 549688705022,
- 549755289600,
- 549755813783,
- 549755813838,
- 549755813854,
- 549755813855,
- 549755813866,
- 549755813869,
- 549755813878,
- 549755813882,
- 549755813886,
- 1099511627566,
- 1099511627710,
- 1099511627738,
- 1099511627759,
- 1099511627761,
- 1099511627774,
- 1425929142271,
- 2199023255518,
- 2199023255522,
- 2199023255543,
- 2199023255550,
- 4363955208191,
- 4398046510080,
- 4398046511079,
- 4398046511086,
- 4398046511095,
- 4398046511099,
- 4398046511102,
- 8727910416382,
- 8791798053935,
- 8796090925055,
- 8796093021443,
- 8796093022019,
- 8796093022158,
- 8796093022179,
- 8796093022183,
- 8796093022189,
- 8796093022190,
- 8796093022193,
- 8796093022198,
- 8796093022205,
- 8796093022206,
- 17583596107870,
- 17592181850110,
- 17592186042886,
- 17592186044038,
- 17592186044358,
- 17592186044366,
- 17592186044378,
- 17592186044399,
- 17592186044410,
- 17592186044411,
- 17592186044413,
- 17592186044414,
- 35184372088715,
- 35184372088768,
- 35184372088798,
- 35184372088822,
- 35184372088826,
- 35184372088829,
- 35184372088830,
- 70368735789055,
- 70368744177430,
- 70368744177637,
- 70368744177647,
- 70368744177651,
- 70368744177655,
- 70368744177658,
- 70368744177659,
- 70368744177660,
- 70368744177661,
- 70368744177662,
- 140737471578110,
- 140737479966719,
- 140737488354759,
- 140737488354847,
- 140737488355141,
- 140737488355274,
- 140737488355294,
- 140737488355301,
- 140737488355302,
- 140737488355303,
- 140737488355310,
- 140737488355313,
- 140737488355318,
- 140737488355319,
- 140737488355326,
- 194613558116351,
- 281453501874175,
- 281470681743359,
- 281474959933438,
- 281474976645119,
- 281474976709518,
- 281474976709694,
- 281474976710235,
- 281474976710282,
- 281474976710339,
- 281474976710469,
- 281474976710551,
- 281474976710602,
- 281474976710606,
- 281474976710625,
- 281474976710626,
- 281474976710631,
- 281474976710637,
- 281474976710638,
- 281474976710639,
- 281474976710645,
- 281474976710647,
- 281474976710653,
- 281474976710654,
- 389227116232702,
- 562907003748350,
- 562949953290238,
- 562949953420470,
- 562949953420678,
- 562949953420938,
- 562949953421102,
- 562949953421250,
- 562949953421262,
- 562949953421274,
- 562949953421278,
- 562949953421279,
- 562949953421290,
- 562949953421291,
- 562949953421293,
- 562949953421294,
- 562949953421295,
- 562949953421297,
- 562949953421303,
- 562949953421306,
- 562949953421308,
- 562949953421310,
- 1125899873288191,
- 1125899906842558,
- 1125899906842582,
- 1125899906842586,
- 1125899906842590,
- 1125899906842593,
- 1125899906842594,
- 1125899906842606,
- 1125899906842607,
- 1125899906842619,
- 1125899906842621,
- 1125899906842622,
- 1460151441686527,
- 2211942517178367,
- 2234929182146559,
- 2248776156708863,
- 2251799813160960,
- 2251799813684483,
- 2251799813685186,
- 2251799813685210,
- 2251799813685214,
- 2251799813685217,
- 2251799813685229,
- 2251799813685231,
- 2251799813685238,
- 2251799813685239,
- 2251799813685242,
- 2251799813685245,
- 2251799813685246,
- 2920302883373054,
- 4423885034356734,
- 4469858364293118,
- 4497552313417726,
- 4503595332402223,
- 4503599627368966,
- 4503599627369927,
- 4503599627370015,
- 4503599627370307,
- 4503599627370309,
- 4503599627370434,
- 4503599627370458,
- 4503599627370462,
- 4503599627370475,
- 4503599627370478,
- 4503599627370479,
- 4503599627370490,
- 4503599627370491,
- 4503599627370494,
- 9007190664804446,
- 9007199187632127,
- 9007199254739854,
- 9007199254740030,
- 9007199254740614,
- 9007199254740618,
- 9007199254740950,
- 9007199254740958,
- 9007199254740963,
- 9007199254740967,
- 9007199254740977,
- 9007199254740982,
- 9007199254740988,
- 9007199254740990,
- 18014398509481926,
- 18014398509481934,
- 18014398509481954,
- 18014398509481975,
- 18014398509481981,
- 18014398509481982,
- 36028797018963547,
- 36028797018963651,
- 36028797018963781,
- 36028797018963863,
- 36028797018963937,
- 36028797018963943,
- 36028797018963947,
- 36028797018963949,
- 36028797018963950,
- 36028797018963951,
- 36028797018963962,
- 36028797018963964,
- 36028797018963966,
- 72056494526300160,
- 72057594037927094,
- 72057594037927302,
- 72057594037927562,
- 72057594037927726,
- 72057594037927819,
- 72057594037927874,
- 72057594037927886,
- 72057594037927894,
- 72057594037927898,
- 72057594037927902,
- 72057594037927915,
- 72057594037927919,
- 72057594037927931,
- 72057594037927933,
- 72057594037927934,
- 144115188075855638,
- 144115188075855830,
- 144115188075855838,
- 144115188075855853,
- 144115188075855857,
- 144115188075855862,
- 144115188075855863,
- 144115188075855866,
- 144115188075855867,
- 144115188075855868,
- 144115188075855869,
- 144115188075855870,
- 288230376151711706,
- 288230376151711713,
- 288230376151711714,
- 288230376151711717,
- 288230376151711726,
- 288230376151711727,
- 288230376151711734,
- 288230376151711738,
- 288230376151711740,
- 288230376151711741,
- 288230376151711742,
- 576460752303423426,
- 576460752303423434,
- 576460752303423454,
- 576460752303423467,
- 576460752303423469,
- 576460752303423471,
- 576460752303423473,
- 576460752303423482,
- 576460752303423486,
- 1117984489315730401,
- 1152921504606846934,
- 1152921504606846938,
- 1152921504606846942,
- 1152921504606846946,
- 1152921504606846974,
- 2305843009213693948,
- 2305843009213693950,
- 3353953467947191203,
- 3816567739388183093,
- 4530058275181297663,
- 4575938696385134591,
- 5675921253449092805,
- 9564978408590137875,
- 9708812670373448219,
- 9963214713607832691,
- 10248191152060862009,
- 10330176681277348905,
- 10365313336655843289,
- 11907422100489763477,
- 12273715991527008853,
- 12297829382473034411,
- 12754194144713244671,
- 14757395258967641293,
- 14897608040525528621,
- 14933078535860113213,
- 14978125529935106013,
- 15580212934572586289,
- 17050145153302519317,
- 17216961135462248175,
- 17256631552825064415,
- 17361641481138401521,
- 18308539860144619519,
- 18421974275759013887,
- 18445336698825998335,
- 18446462598732840959,
- 18446726481523507199,
- 18446744065119617023,
- 18446744069414583343,
- 18446744069414584319,
- 18446744069414584320,
- 18446744069414584321,
- 18446744073709486079,
- 18446744073709550851,
- 18446744073709551047,
- 18446744073709551135,
- 18446744073709551195,
- 18446744073709551299,
- 18446744073709551427,
- 18446744073709551429,
- 18446744073709551499,
- 18446744073709551511,
- 18446744073709551583,
- 18446744073709551585,
- 18446744073709551587,
- 18446744073709551589,
- 18446744073709551591,
- 18446744073709551595,
- 18446744073709551597,
- 18446744073709551599,
- 18446744073709551601,
- 18446744073709551603,
- 18446744073709551605,
- 18446744073709551607,
- 18446744073709551611,
- 18446744073709551613,
- 18446744073709551614,
- 77371252455336267181195254,
- 77371252455336267181195262,
- 340282366841710300967557013911933812736,
-])))
-
-def log2_up(i):
- return int(math.ceil(math.log(i, 2)))
-
-def word(i, width=None):
- assert(i >= 0)
- if width is None:
- if i == 0: return word(i, width=1)
- return word(i, width=2**log2_up(log2_up(i + 1)))
- return 'WO~' + '~'.join(bin(i)[2:].rjust(width, '0'))
-
-word_formats = tuple(sorted([(n, hex(n), bin(n), word(n)) for n in nums]))
-
-def header():
- return (r"""Require Import Coq.ZArith.ZArith.
-Require Export Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Export bbv.WordScope.
-Require Export Crypto.Util.Notations.
-
-Notation Const x := (Op (OpConst x) TT).
-(""" + r"""* python:
-<<
-""" +
- open(__file__).read() +
- r""">>
- *""" + r""")""")
-
-
-def hex_nots():
- return ('\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 word_formats))
-def bin_nots():
- return ('\n'.join('''Notation "'%s'" (* %d (%s) *)\n := (Const %s).''' % (b, d, h, w)
- for d, h, b, w in word_formats))
-
-with open('HexNotationConstants.v', 'w') as f:
- f.write(header() + '\n' + hex_nots() + '\n')
-
-with open('BinaryNotationConstants.v', 'w') as f:
- f.write(header() + '\n' + bin_nots() + '\n')
->>
- *)
-Notation "'0x0'" (* 0 (0x0) *)
- := (Const 0%Z).
-Notation "'0x0'" (* 0 (0x0) *)
- := (Const WO~0).
-Notation "'0x1'" (* 1 (0x1) *)
- := (Const 1%Z).
-Notation "'0x1'" (* 1 (0x1) *)
- := (Const WO~1).
-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 "'0x4'" (* 4 (0x4) *)
- := (Const 4%Z).
-Notation "'0x4'" (* 4 (0x4) *)
- := (Const WO~0~1~0~0).
-Notation "'0x5'" (* 5 (0x5) *)
- := (Const 5%Z).
-Notation "'0x5'" (* 5 (0x5) *)
- := (Const WO~0~1~0~1).
-Notation "'0x6'" (* 6 (0x6) *)
- := (Const 6%Z).
-Notation "'0x6'" (* 6 (0x6) *)
- := (Const WO~0~1~1~0).
-Notation "'0x7'" (* 7 (0x7) *)
- := (Const 7%Z).
-Notation "'0x7'" (* 7 (0x7) *)
- := (Const WO~0~1~1~1).
-Notation "'0x8'" (* 8 (0x8) *)
- := (Const 8%Z).
-Notation "'0x8'" (* 8 (0x8) *)
- := (Const WO~1~0~0~0).
-Notation "'0x9'" (* 9 (0x9) *)
- := (Const 9%Z).
-Notation "'0x9'" (* 9 (0x9) *)
- := (Const WO~1~0~0~1).
-Notation "'0xa'" (* 10 (0xa) *)
- := (Const 10%Z).
-Notation "'0xa'" (* 10 (0xa) *)
- := (Const WO~1~0~1~0).
-Notation "'0xb'" (* 11 (0xb) *)
- := (Const 11%Z).
-Notation "'0xb'" (* 11 (0xb) *)
- := (Const WO~1~0~1~1).
-Notation "'0xc'" (* 12 (0xc) *)
- := (Const 12%Z).
-Notation "'0xc'" (* 12 (0xc) *)
- := (Const WO~1~1~0~0).
-Notation "'0xd'" (* 13 (0xd) *)
- := (Const 13%Z).
-Notation "'0xd'" (* 13 (0xd) *)
- := (Const WO~1~1~0~1).
-Notation "'0xe'" (* 14 (0xe) *)
- := (Const 14%Z).
-Notation "'0xe'" (* 14 (0xe) *)
- := (Const WO~1~1~1~0).
-Notation "'0xf'" (* 15 (0xf) *)
- := (Const 15%Z).
-Notation "'0xf'" (* 15 (0xf) *)
- := (Const WO~1~1~1~1).
-Notation "'0x10'" (* 16 (0x10) *)
- := (Const 16%Z).
-Notation "'0x10'" (* 16 (0x10) *)
- := (Const WO~0~0~0~1~0~0~0~0).
-Notation "'0x11'" (* 17 (0x11) *)
- := (Const 17%Z).
-Notation "'0x11'" (* 17 (0x11) *)
- := (Const WO~0~0~0~1~0~0~0~1).
-Notation "'0x12'" (* 18 (0x12) *)
- := (Const 18%Z).
-Notation "'0x12'" (* 18 (0x12) *)
- := (Const WO~0~0~0~1~0~0~1~0).
-Notation "'0x13'" (* 19 (0x13) *)
- := (Const 19%Z).
-Notation "'0x13'" (* 19 (0x13) *)
- := (Const WO~0~0~0~1~0~0~1~1).
-Notation "'0x14'" (* 20 (0x14) *)
- := (Const 20%Z).
-Notation "'0x14'" (* 20 (0x14) *)
- := (Const WO~0~0~0~1~0~1~0~0).
-Notation "'0x15'" (* 21 (0x15) *)
- := (Const 21%Z).
-Notation "'0x15'" (* 21 (0x15) *)
- := (Const WO~0~0~0~1~0~1~0~1).
-Notation "'0x16'" (* 22 (0x16) *)
- := (Const 22%Z).
-Notation "'0x16'" (* 22 (0x16) *)
- := (Const WO~0~0~0~1~0~1~1~0).
-Notation "'0x17'" (* 23 (0x17) *)
- := (Const 23%Z).
-Notation "'0x17'" (* 23 (0x17) *)
- := (Const WO~0~0~0~1~0~1~1~1).
-Notation "'0x18'" (* 24 (0x18) *)
- := (Const 24%Z).
-Notation "'0x18'" (* 24 (0x18) *)
- := (Const WO~0~0~0~1~1~0~0~0).
-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 "'0x1d'" (* 29 (0x1d) *)
- := (Const 29%Z).
-Notation "'0x1d'" (* 29 (0x1d) *)
- := (Const WO~0~0~0~1~1~1~0~1).
-Notation "'0x1e'" (* 30 (0x1e) *)
- := (Const 30%Z).
-Notation "'0x1e'" (* 30 (0x1e) *)
- := (Const WO~0~0~0~1~1~1~1~0).
-Notation "'0x1f'" (* 31 (0x1f) *)
- := (Const 31%Z).
-Notation "'0x1f'" (* 31 (0x1f) *)
- := (Const WO~0~0~0~1~1~1~1~1).
-Notation "'0x20'" (* 32 (0x20) *)
- := (Const 32%Z).
-Notation "'0x20'" (* 32 (0x20) *)
- := (Const WO~0~0~1~0~0~0~0~0).
-Notation "'0x21'" (* 33 (0x21) *)
- := (Const 33%Z).
-Notation "'0x21'" (* 33 (0x21) *)
- := (Const WO~0~0~1~0~0~0~0~1).
-Notation "'0x22'" (* 34 (0x22) *)
- := (Const 34%Z).
-Notation "'0x22'" (* 34 (0x22) *)
- := (Const WO~0~0~1~0~0~0~1~0).
-Notation "'0x23'" (* 35 (0x23) *)
- := (Const 35%Z).
-Notation "'0x23'" (* 35 (0x23) *)
- := (Const WO~0~0~1~0~0~0~1~1).
-Notation "'0x24'" (* 36 (0x24) *)
- := (Const 36%Z).
-Notation "'0x24'" (* 36 (0x24) *)
- := (Const WO~0~0~1~0~0~1~0~0).
-Notation "'0x25'" (* 37 (0x25) *)
- := (Const 37%Z).
-Notation "'0x25'" (* 37 (0x25) *)
- := (Const WO~0~0~1~0~0~1~0~1).
-Notation "'0x26'" (* 38 (0x26) *)
- := (Const 38%Z).
-Notation "'0x26'" (* 38 (0x26) *)
- := (Const WO~0~0~1~0~0~1~1~0).
-Notation "'0x27'" (* 39 (0x27) *)
- := (Const 39%Z).
-Notation "'0x27'" (* 39 (0x27) *)
- := (Const WO~0~0~1~0~0~1~1~1).
-Notation "'0x28'" (* 40 (0x28) *)
- := (Const 40%Z).
-Notation "'0x28'" (* 40 (0x28) *)
- := (Const WO~0~0~1~0~1~0~0~0).
-Notation "'0x29'" (* 41 (0x29) *)
- := (Const 41%Z).
-Notation "'0x29'" (* 41 (0x29) *)
- := (Const WO~0~0~1~0~1~0~0~1).
-Notation "'0x2a'" (* 42 (0x2a) *)
- := (Const 42%Z).
-Notation "'0x2a'" (* 42 (0x2a) *)
- := (Const WO~0~0~1~0~1~0~1~0).
-Notation "'0x2b'" (* 43 (0x2b) *)
- := (Const 43%Z).
-Notation "'0x2b'" (* 43 (0x2b) *)
- := (Const WO~0~0~1~0~1~0~1~1).
-Notation "'0x2c'" (* 44 (0x2c) *)
- := (Const 44%Z).
-Notation "'0x2c'" (* 44 (0x2c) *)
- := (Const WO~0~0~1~0~1~1~0~0).
-Notation "'0x2d'" (* 45 (0x2d) *)
- := (Const 45%Z).
-Notation "'0x2d'" (* 45 (0x2d) *)
- := (Const WO~0~0~1~0~1~1~0~1).
-Notation "'0x2e'" (* 46 (0x2e) *)
- := (Const 46%Z).
-Notation "'0x2e'" (* 46 (0x2e) *)
- := (Const WO~0~0~1~0~1~1~1~0).
-Notation "'0x2f'" (* 47 (0x2f) *)
- := (Const 47%Z).
-Notation "'0x2f'" (* 47 (0x2f) *)
- := (Const WO~0~0~1~0~1~1~1~1).
-Notation "'0x30'" (* 48 (0x30) *)
- := (Const 48%Z).
-Notation "'0x30'" (* 48 (0x30) *)
- := (Const WO~0~0~1~1~0~0~0~0).
-Notation "'0x31'" (* 49 (0x31) *)
- := (Const 49%Z).
-Notation "'0x31'" (* 49 (0x31) *)
- := (Const WO~0~0~1~1~0~0~0~1).
-Notation "'0x32'" (* 50 (0x32) *)
- := (Const 50%Z).
-Notation "'0x32'" (* 50 (0x32) *)
- := (Const WO~0~0~1~1~0~0~1~0).
-Notation "'0x33'" (* 51 (0x33) *)
- := (Const 51%Z).
-Notation "'0x33'" (* 51 (0x33) *)
- := (Const WO~0~0~1~1~0~0~1~1).
-Notation "'0x34'" (* 52 (0x34) *)
- := (Const 52%Z).
-Notation "'0x34'" (* 52 (0x34) *)
- := (Const WO~0~0~1~1~0~1~0~0).
-Notation "'0x35'" (* 53 (0x35) *)
- := (Const 53%Z).
-Notation "'0x35'" (* 53 (0x35) *)
- := (Const WO~0~0~1~1~0~1~0~1).
-Notation "'0x36'" (* 54 (0x36) *)
- := (Const 54%Z).
-Notation "'0x36'" (* 54 (0x36) *)
- := (Const WO~0~0~1~1~0~1~1~0).
-Notation "'0x37'" (* 55 (0x37) *)
- := (Const 55%Z).
-Notation "'0x37'" (* 55 (0x37) *)
- := (Const WO~0~0~1~1~0~1~1~1).
-Notation "'0x38'" (* 56 (0x38) *)
- := (Const 56%Z).
-Notation "'0x38'" (* 56 (0x38) *)
- := (Const WO~0~0~1~1~1~0~0~0).
-Notation "'0x39'" (* 57 (0x39) *)
- := (Const 57%Z).
-Notation "'0x39'" (* 57 (0x39) *)
- := (Const WO~0~0~1~1~1~0~0~1).
-Notation "'0x3a'" (* 58 (0x3a) *)
- := (Const 58%Z).
-Notation "'0x3a'" (* 58 (0x3a) *)
- := (Const WO~0~0~1~1~1~0~1~0).
-Notation "'0x3b'" (* 59 (0x3b) *)
- := (Const 59%Z).
-Notation "'0x3b'" (* 59 (0x3b) *)
- := (Const WO~0~0~1~1~1~0~1~1).
-Notation "'0x3c'" (* 60 (0x3c) *)
- := (Const 60%Z).
-Notation "'0x3c'" (* 60 (0x3c) *)
- := (Const WO~0~0~1~1~1~1~0~0).
-Notation "'0x3d'" (* 61 (0x3d) *)
- := (Const 61%Z).
-Notation "'0x3d'" (* 61 (0x3d) *)
- := (Const WO~0~0~1~1~1~1~0~1).
-Notation "'0x3e'" (* 62 (0x3e) *)
- := (Const 62%Z).
-Notation "'0x3e'" (* 62 (0x3e) *)
- := (Const WO~0~0~1~1~1~1~1~0).
-Notation "'0x3f'" (* 63 (0x3f) *)
- := (Const 63%Z).
-Notation "'0x3f'" (* 63 (0x3f) *)
- := (Const WO~0~0~1~1~1~1~1~1).
-Notation "'0x40'" (* 64 (0x40) *)
- := (Const 64%Z).
-Notation "'0x40'" (* 64 (0x40) *)
- := (Const WO~0~1~0~0~0~0~0~0).
-Notation "'0x41'" (* 65 (0x41) *)
- := (Const 65%Z).
-Notation "'0x41'" (* 65 (0x41) *)
- := (Const WO~0~1~0~0~0~0~0~1).
-Notation "'0x42'" (* 66 (0x42) *)
- := (Const 66%Z).
-Notation "'0x42'" (* 66 (0x42) *)
- := (Const WO~0~1~0~0~0~0~1~0).
-Notation "'0x43'" (* 67 (0x43) *)
- := (Const 67%Z).
-Notation "'0x43'" (* 67 (0x43) *)
- := (Const WO~0~1~0~0~0~0~1~1).
-Notation "'0x44'" (* 68 (0x44) *)
- := (Const 68%Z).
-Notation "'0x44'" (* 68 (0x44) *)
- := (Const WO~0~1~0~0~0~1~0~0).
-Notation "'0x45'" (* 69 (0x45) *)
- := (Const 69%Z).
-Notation "'0x45'" (* 69 (0x45) *)
- := (Const WO~0~1~0~0~0~1~0~1).
-Notation "'0x46'" (* 70 (0x46) *)
- := (Const 70%Z).
-Notation "'0x46'" (* 70 (0x46) *)
- := (Const WO~0~1~0~0~0~1~1~0).
-Notation "'0x47'" (* 71 (0x47) *)
- := (Const 71%Z).
-Notation "'0x47'" (* 71 (0x47) *)
- := (Const WO~0~1~0~0~0~1~1~1).
-Notation "'0x48'" (* 72 (0x48) *)
- := (Const 72%Z).
-Notation "'0x48'" (* 72 (0x48) *)
- := (Const WO~0~1~0~0~1~0~0~0).
-Notation "'0x49'" (* 73 (0x49) *)
- := (Const 73%Z).
-Notation "'0x49'" (* 73 (0x49) *)
- := (Const WO~0~1~0~0~1~0~0~1).
-Notation "'0x4a'" (* 74 (0x4a) *)
- := (Const 74%Z).
-Notation "'0x4a'" (* 74 (0x4a) *)
- := (Const WO~0~1~0~0~1~0~1~0).
-Notation "'0x4b'" (* 75 (0x4b) *)
- := (Const 75%Z).
-Notation "'0x4b'" (* 75 (0x4b) *)
- := (Const WO~0~1~0~0~1~0~1~1).
-Notation "'0x4c'" (* 76 (0x4c) *)
- := (Const 76%Z).
-Notation "'0x4c'" (* 76 (0x4c) *)
- := (Const WO~0~1~0~0~1~1~0~0).
-Notation "'0x4d'" (* 77 (0x4d) *)
- := (Const 77%Z).
-Notation "'0x4d'" (* 77 (0x4d) *)
- := (Const WO~0~1~0~0~1~1~0~1).
-Notation "'0x4e'" (* 78 (0x4e) *)
- := (Const 78%Z).
-Notation "'0x4e'" (* 78 (0x4e) *)
- := (Const WO~0~1~0~0~1~1~1~0).
-Notation "'0x4f'" (* 79 (0x4f) *)
- := (Const 79%Z).
-Notation "'0x4f'" (* 79 (0x4f) *)
- := (Const WO~0~1~0~0~1~1~1~1).
-Notation "'0x50'" (* 80 (0x50) *)
- := (Const 80%Z).
-Notation "'0x50'" (* 80 (0x50) *)
- := (Const WO~0~1~0~1~0~0~0~0).
-Notation "'0x51'" (* 81 (0x51) *)
- := (Const 81%Z).
-Notation "'0x51'" (* 81 (0x51) *)
- := (Const WO~0~1~0~1~0~0~0~1).
-Notation "'0x52'" (* 82 (0x52) *)
- := (Const 82%Z).
-Notation "'0x52'" (* 82 (0x52) *)
- := (Const WO~0~1~0~1~0~0~1~0).
-Notation "'0x53'" (* 83 (0x53) *)
- := (Const 83%Z).
-Notation "'0x53'" (* 83 (0x53) *)
- := (Const WO~0~1~0~1~0~0~1~1).
-Notation "'0x54'" (* 84 (0x54) *)
- := (Const 84%Z).
-Notation "'0x54'" (* 84 (0x54) *)
- := (Const WO~0~1~0~1~0~1~0~0).
-Notation "'0x55'" (* 85 (0x55) *)
- := (Const 85%Z).
-Notation "'0x55'" (* 85 (0x55) *)
- := (Const WO~0~1~0~1~0~1~0~1).
-Notation "'0x56'" (* 86 (0x56) *)
- := (Const 86%Z).
-Notation "'0x56'" (* 86 (0x56) *)
- := (Const WO~0~1~0~1~0~1~1~0).
-Notation "'0x57'" (* 87 (0x57) *)
- := (Const 87%Z).
-Notation "'0x57'" (* 87 (0x57) *)
- := (Const WO~0~1~0~1~0~1~1~1).
-Notation "'0x58'" (* 88 (0x58) *)
- := (Const 88%Z).
-Notation "'0x58'" (* 88 (0x58) *)
- := (Const WO~0~1~0~1~1~0~0~0).
-Notation "'0x59'" (* 89 (0x59) *)
- := (Const 89%Z).
-Notation "'0x59'" (* 89 (0x59) *)
- := (Const WO~0~1~0~1~1~0~0~1).
-Notation "'0x5a'" (* 90 (0x5a) *)
- := (Const 90%Z).
-Notation "'0x5a'" (* 90 (0x5a) *)
- := (Const WO~0~1~0~1~1~0~1~0).
-Notation "'0x5b'" (* 91 (0x5b) *)
- := (Const 91%Z).
-Notation "'0x5b'" (* 91 (0x5b) *)
- := (Const WO~0~1~0~1~1~0~1~1).
-Notation "'0x5c'" (* 92 (0x5c) *)
- := (Const 92%Z).
-Notation "'0x5c'" (* 92 (0x5c) *)
- := (Const WO~0~1~0~1~1~1~0~0).
-Notation "'0x5d'" (* 93 (0x5d) *)
- := (Const 93%Z).
-Notation "'0x5d'" (* 93 (0x5d) *)
- := (Const WO~0~1~0~1~1~1~0~1).
-Notation "'0x5e'" (* 94 (0x5e) *)
- := (Const 94%Z).
-Notation "'0x5e'" (* 94 (0x5e) *)
- := (Const WO~0~1~0~1~1~1~1~0).
-Notation "'0x5f'" (* 95 (0x5f) *)
- := (Const 95%Z).
-Notation "'0x5f'" (* 95 (0x5f) *)
- := (Const WO~0~1~0~1~1~1~1~1).
-Notation "'0x60'" (* 96 (0x60) *)
- := (Const 96%Z).
-Notation "'0x60'" (* 96 (0x60) *)
- := (Const WO~0~1~1~0~0~0~0~0).
-Notation "'0x61'" (* 97 (0x61) *)
- := (Const 97%Z).
-Notation "'0x61'" (* 97 (0x61) *)
- := (Const WO~0~1~1~0~0~0~0~1).
-Notation "'0x62'" (* 98 (0x62) *)
- := (Const 98%Z).
-Notation "'0x62'" (* 98 (0x62) *)
- := (Const WO~0~1~1~0~0~0~1~0).
-Notation "'0x63'" (* 99 (0x63) *)
- := (Const 99%Z).
-Notation "'0x63'" (* 99 (0x63) *)
- := (Const WO~0~1~1~0~0~0~1~1).
-Notation "'0x64'" (* 100 (0x64) *)
- := (Const 100%Z).
-Notation "'0x64'" (* 100 (0x64) *)
- := (Const WO~0~1~1~0~0~1~0~0).
-Notation "'0x65'" (* 101 (0x65) *)
- := (Const 101%Z).
-Notation "'0x65'" (* 101 (0x65) *)
- := (Const WO~0~1~1~0~0~1~0~1).
-Notation "'0x66'" (* 102 (0x66) *)
- := (Const 102%Z).
-Notation "'0x66'" (* 102 (0x66) *)
- := (Const WO~0~1~1~0~0~1~1~0).
-Notation "'0x67'" (* 103 (0x67) *)
- := (Const 103%Z).
-Notation "'0x67'" (* 103 (0x67) *)
- := (Const WO~0~1~1~0~0~1~1~1).
-Notation "'0x68'" (* 104 (0x68) *)
- := (Const 104%Z).
-Notation "'0x68'" (* 104 (0x68) *)
- := (Const WO~0~1~1~0~1~0~0~0).
-Notation "'0x69'" (* 105 (0x69) *)
- := (Const 105%Z).
-Notation "'0x69'" (* 105 (0x69) *)
- := (Const WO~0~1~1~0~1~0~0~1).
-Notation "'0x6a'" (* 106 (0x6a) *)
- := (Const 106%Z).
-Notation "'0x6a'" (* 106 (0x6a) *)
- := (Const WO~0~1~1~0~1~0~1~0).
-Notation "'0x6b'" (* 107 (0x6b) *)
- := (Const 107%Z).
-Notation "'0x6b'" (* 107 (0x6b) *)
- := (Const WO~0~1~1~0~1~0~1~1).
-Notation "'0x6c'" (* 108 (0x6c) *)
- := (Const 108%Z).
-Notation "'0x6c'" (* 108 (0x6c) *)
- := (Const WO~0~1~1~0~1~1~0~0).
-Notation "'0x6d'" (* 109 (0x6d) *)
- := (Const 109%Z).
-Notation "'0x6d'" (* 109 (0x6d) *)
- := (Const WO~0~1~1~0~1~1~0~1).
-Notation "'0x6e'" (* 110 (0x6e) *)
- := (Const 110%Z).
-Notation "'0x6e'" (* 110 (0x6e) *)
- := (Const WO~0~1~1~0~1~1~1~0).
-Notation "'0x6f'" (* 111 (0x6f) *)
- := (Const 111%Z).
-Notation "'0x6f'" (* 111 (0x6f) *)
- := (Const WO~0~1~1~0~1~1~1~1).
-Notation "'0x70'" (* 112 (0x70) *)
- := (Const 112%Z).
-Notation "'0x70'" (* 112 (0x70) *)
- := (Const WO~0~1~1~1~0~0~0~0).
-Notation "'0x71'" (* 113 (0x71) *)
- := (Const 113%Z).
-Notation "'0x71'" (* 113 (0x71) *)
- := (Const WO~0~1~1~1~0~0~0~1).
-Notation "'0x72'" (* 114 (0x72) *)
- := (Const 114%Z).
-Notation "'0x72'" (* 114 (0x72) *)
- := (Const WO~0~1~1~1~0~0~1~0).
-Notation "'0x73'" (* 115 (0x73) *)
- := (Const 115%Z).
-Notation "'0x73'" (* 115 (0x73) *)
- := (Const WO~0~1~1~1~0~0~1~1).
-Notation "'0x74'" (* 116 (0x74) *)
- := (Const 116%Z).
-Notation "'0x74'" (* 116 (0x74) *)
- := (Const WO~0~1~1~1~0~1~0~0).
-Notation "'0x75'" (* 117 (0x75) *)
- := (Const 117%Z).
-Notation "'0x75'" (* 117 (0x75) *)
- := (Const WO~0~1~1~1~0~1~0~1).
-Notation "'0x76'" (* 118 (0x76) *)
- := (Const 118%Z).
-Notation "'0x76'" (* 118 (0x76) *)
- := (Const WO~0~1~1~1~0~1~1~0).
-Notation "'0x77'" (* 119 (0x77) *)
- := (Const 119%Z).
-Notation "'0x77'" (* 119 (0x77) *)
- := (Const WO~0~1~1~1~0~1~1~1).
-Notation "'0x78'" (* 120 (0x78) *)
- := (Const 120%Z).
-Notation "'0x78'" (* 120 (0x78) *)
- := (Const WO~0~1~1~1~1~0~0~0).
-Notation "'0x79'" (* 121 (0x79) *)
- := (Const 121%Z).
-Notation "'0x79'" (* 121 (0x79) *)
- := (Const WO~0~1~1~1~1~0~0~1).
-Notation "'0x7a'" (* 122 (0x7a) *)
- := (Const 122%Z).
-Notation "'0x7a'" (* 122 (0x7a) *)
- := (Const WO~0~1~1~1~1~0~1~0).
-Notation "'0x7b'" (* 123 (0x7b) *)
- := (Const 123%Z).
-Notation "'0x7b'" (* 123 (0x7b) *)
- := (Const WO~0~1~1~1~1~0~1~1).
-Notation "'0x7c'" (* 124 (0x7c) *)
- := (Const 124%Z).
-Notation "'0x7c'" (* 124 (0x7c) *)
- := (Const WO~0~1~1~1~1~1~0~0).
-Notation "'0x7d'" (* 125 (0x7d) *)
- := (Const 125%Z).
-Notation "'0x7d'" (* 125 (0x7d) *)
- := (Const WO~0~1~1~1~1~1~0~1).
-Notation "'0x7e'" (* 126 (0x7e) *)
- := (Const 126%Z).
-Notation "'0x7e'" (* 126 (0x7e) *)
- := (Const WO~0~1~1~1~1~1~1~0).
-Notation "'0x7f'" (* 127 (0x7f) *)
- := (Const 127%Z).
-Notation "'0x7f'" (* 127 (0x7f) *)
- := (Const WO~0~1~1~1~1~1~1~1).
-Notation "'0x80'" (* 128 (0x80) *)
- := (Const 128%Z).
-Notation "'0x80'" (* 128 (0x80) *)
- := (Const WO~1~0~0~0~0~0~0~0).
-Notation "'0x81'" (* 129 (0x81) *)
- := (Const 129%Z).
-Notation "'0x81'" (* 129 (0x81) *)
- := (Const WO~1~0~0~0~0~0~0~1).
-Notation "'0xbb'" (* 187 (0xbb) *)
- := (Const 187%Z).
-Notation "'0xbb'" (* 187 (0xbb) *)
- := (Const WO~1~0~1~1~1~0~1~1).
-Notation "'0xbd'" (* 189 (0xbd) *)
- := (Const 189%Z).
-Notation "'0xbd'" (* 189 (0xbd) *)
- := (Const WO~1~0~1~1~1~1~0~1).
-Notation "'0xff'" (* 255 (0xff) *)
- := (Const 255%Z).
-Notation "'0xff'" (* 255 (0xff) *)
- := (Const WO~1~1~1~1~1~1~1~1).
-Notation "'0x100'" (* 256 (0x100) *)
- := (Const 256%Z).
-Notation "'0x100'" (* 256 (0x100) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0).
-Notation "'0x101'" (* 257 (0x101) *)
- := (Const 257%Z).
-Notation "'0x101'" (* 257 (0x101) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~1).
-Notation "'0x13d'" (* 317 (0x13d) *)
- := (Const 317%Z).
-Notation "'0x13d'" (* 317 (0x13d) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~1~1~1~1~0~1).
-Notation "'0x1a5'" (* 421 (0x1a5) *)
- := (Const 421%Z).
-Notation "'0x1a5'" (* 421 (0x1a5) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~0~1~0~0~1~0~1).
-Notation "'0x1e1'" (* 481 (0x1e1) *)
- := (Const 481%Z).
-Notation "'0x1e1'" (* 481 (0x1e1) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~0~0~0~0~1).
-Notation "'0x1ff'" (* 511 (0x1ff) *)
- := (Const 511%Z).
-Notation "'0x1ff'" (* 511 (0x1ff) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1).
-Notation "'0x200'" (* 512 (0x200) *)
- := (Const 512%Z).
-Notation "'0x200'" (* 512 (0x200) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0).
-Notation "'0x201'" (* 513 (0x201) *)
- := (Const 513%Z).
-Notation "'0x201'" (* 513 (0x201) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~1).
-Notation "'0x239'" (* 569 (0x239) *)
- := (Const 569%Z).
-Notation "'0x239'" (* 569 (0x239) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~1~1~1~0~0~1).
-Notation "'0x2fd'" (* 765 (0x2fd) *)
- := (Const 765%Z).
-Notation "'0x2fd'" (* 765 (0x2fd) *)
- := (Const WO~0~0~0~0~0~0~1~0~1~1~1~1~1~1~0~1).
-Notation "'0x3d1'" (* 977 (0x3d1) *)
- := (Const 977%Z).
-Notation "'0x3d1'" (* 977 (0x3d1) *)
- := (Const WO~0~0~0~0~0~0~1~1~1~1~0~1~0~0~0~1).
-Notation "'0x3ff'" (* 1023 (0x3ff) *)
- := (Const 1023%Z).
-Notation "'0x3ff'" (* 1023 (0x3ff) *)
- := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x400'" (* 1024 (0x400) *)
- := (Const 1024%Z).
-Notation "'0x400'" (* 1024 (0x400) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x401'" (* 1025 (0x401) *)
- := (Const 1025%Z).
-Notation "'0x401'" (* 1025 (0x401) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ff'" (* 2047 (0x7ff) *)
- := (Const 2047%Z).
-Notation "'0x7ff'" (* 2047 (0x7ff) *)
- := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x800'" (* 2048 (0x800) *)
- := (Const 2048%Z).
-Notation "'0x800'" (* 2048 (0x800) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x801'" (* 2049 (0x801) *)
- := (Const 2049%Z).
-Notation "'0x801'" (* 2049 (0x801) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfff'" (* 4095 (0xfff) *)
- := (Const 4095%Z).
-Notation "'0xfff'" (* 4095 (0xfff) *)
- := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x1000'" (* 4096 (0x1000) *)
- := (Const 4096%Z).
-Notation "'0x1000'" (* 4096 (0x1000) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1001'" (* 4097 (0x1001) *)
- := (Const 4097%Z).
-Notation "'0x1001'" (* 4097 (0x1001) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x14bf'" (* 5311 (0x14bf) *)
- := (Const 5311%Z).
-Notation "'0x14bf'" (* 5311 (0x14bf) *)
- := (Const WO~0~0~0~1~0~1~0~0~1~0~1~1~1~1~1~1).
-Notation "'0x1fff'" (* 8191 (0x1fff) *)
- := (Const 8191%Z).
-Notation "'0x1fff'" (* 8191 (0x1fff) *)
- := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x2000'" (* 8192 (0x2000) *)
- := (Const 8192%Z).
-Notation "'0x2000'" (* 8192 (0x2000) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2001'" (* 8193 (0x2001) *)
- := (Const 8193%Z).
-Notation "'0x2001'" (* 8193 (0x2001) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fff'" (* 16383 (0x3fff) *)
- := (Const 16383%Z).
-Notation "'0x3fff'" (* 16383 (0x3fff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x4000'" (* 16384 (0x4000) *)
- := (Const 16384%Z).
-Notation "'0x4000'" (* 16384 (0x4000) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4001'" (* 16385 (0x4001) *)
- := (Const 16385%Z).
-Notation "'0x4001'" (* 16385 (0x4001) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffd'" (* 32765 (0x7ffd) *)
- := (Const 32765%Z).
-Notation "'0x7ffd'" (* 32765 (0x7ffd) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1).
-Notation "'0x7ffe'" (* 32766 (0x7ffe) *)
- := (Const 32766%Z).
-Notation "'0x7ffe'" (* 32766 (0x7ffe) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x7fff'" (* 32767 (0x7fff) *)
- := (Const 32767%Z).
-Notation "'0x7fff'" (* 32767 (0x7fff) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x8000'" (* 32768 (0x8000) *)
- := (Const 32768%Z).
-Notation "'0x8000'" (* 32768 (0x8000) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8001'" (* 32769 (0x8001) *)
- := (Const 32769%Z).
-Notation "'0x8001'" (* 32769 (0x8001) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffa'" (* 65530 (0xfffa) *)
- := (Const 65530%Z).
-Notation "'0xfffa'" (* 65530 (0xfffa) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0).
-Notation "'0xfffb'" (* 65531 (0xfffb) *)
- := (Const 65531%Z).
-Notation "'0xfffb'" (* 65531 (0xfffb) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1).
-Notation "'0xfffe'" (* 65534 (0xfffe) *)
- := (Const 65534%Z).
-Notation "'0xfffe'" (* 65534 (0xfffe) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0xffff'" (* 65535 (0xffff) *)
- := (Const 65535%Z).
-Notation "'0xffff'" (* 65535 (0xffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x10000'" (* 65536 (0x10000) *)
- := (Const 65536%Z).
-Notation "'0x10000'" (* 65536 (0x10000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10001'" (* 65537 (0x10001) *)
- := (Const 65537%Z).
-Notation "'0x10001'" (* 65537 (0x10001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1bfff'" (* 114687 (0x1bfff) *)
- := (Const 114687%Z).
-Notation "'0x1bfff'" (* 114687 (0x1bfff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x1db41'" (* 121665 (0x1db41) *)
- := (Const 121665%Z).
-Notation "'0x1db41'" (* 121665 (0x1db41) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~0~1~1~0~1~1~0~1~0~0~0~0~0~1).
-Notation "'0x1fd03'" (* 130307 (0x1fd03) *)
- := (Const 130307%Z).
-Notation "'0x1fd03'" (* 130307 (0x1fd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1).
-Notation "'0x1ffdf'" (* 131039 (0x1ffdf) *)
- := (Const 131039%Z).
-Notation "'0x1ffdf'" (* 131039 (0x1ffdf) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1).
-Notation "'0x1fff6'" (* 131062 (0x1fff6) *)
- := (Const 131062%Z).
-Notation "'0x1fff6'" (* 131062 (0x1fff6) *)
- := (Const WO~0~0~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~0~1~1~0).
-Notation "'0x1fffb'" (* 131067 (0x1fffb) *)
- := (Const 131067%Z).
-Notation "'0x1fffb'" (* 131067 (0x1fffb) *)
- := (Const WO~0~0~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~0~1~1).
-Notation "'0x1ffff'" (* 131071 (0x1ffff) *)
- := (Const 131071%Z).
-Notation "'0x1ffff'" (* 131071 (0x1ffff) *)
- := (Const WO~0~0~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).
-Notation "'0x20000'" (* 131072 (0x20000) *)
- := (Const 131072%Z).
-Notation "'0x20000'" (* 131072 (0x20000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20001'" (* 131073 (0x20001) *)
- := (Const 131073%Z).
-Notation "'0x20001'" (* 131073 (0x20001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fc2f'" (* 261167 (0x3fc2f) *)
- := (Const 261167%Z).
-Notation "'0x3fc2f'" (* 261167 (0x3fc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0x3fdc7'" (* 261575 (0x3fdc7) *)
- := (Const 261575%Z).
-Notation "'0x3fdc7'" (* 261575 (0x3fdc7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1).
-Notation "'0x3fff3'" (* 262131 (0x3fff3) *)
- := (Const 262131%Z).
-Notation "'0x3fff3'" (* 262131 (0x3fff3) *)
- := (Const WO~0~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~0~0~1~1).
-Notation "'0x3fffb'" (* 262139 (0x3fffb) *)
- := (Const 262139%Z).
-Notation "'0x3fffb'" (* 262139 (0x3fffb) *)
- := (Const WO~0~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~0~1~1).
-Notation "'0x3fffe'" (* 262142 (0x3fffe) *)
- := (Const 262142%Z).
-Notation "'0x3fffe'" (* 262142 (0x3fffe) *)
- := (Const WO~0~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~0).
-Notation "'0x3ffff'" (* 262143 (0x3ffff) *)
- := (Const 262143%Z).
-Notation "'0x3ffff'" (* 262143 (0x3ffff) *)
- := (Const WO~0~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).
-Notation "'0x40000'" (* 262144 (0x40000) *)
- := (Const 262144%Z).
-Notation "'0x40000'" (* 262144 (0x40000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40001'" (* 262145 (0x40001) *)
- := (Const 262145%Z).
-Notation "'0x40001'" (* 262145 (0x40001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fe1f'" (* 523807 (0x7fe1f) *)
- := (Const 523807%Z).
-Notation "'0x7fe1f'" (* 523807 (0x7fe1f) *)
- := (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~0~0~0~0~1~1~1~1~1).
-Notation "'0x7ff45'" (* 524101 (0x7ff45) *)
- := (Const 524101%Z).
-Notation "'0x7ff45'" (* 524101 (0x7ff45) *)
- := (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~0~1~0~0~0~1~0~1).
-Notation "'0x7ffbf'" (* 524223 (0x7ffbf) *)
- := (Const 524223%Z).
-Notation "'0x7ffbf'" (* 524223 (0x7ffbf) *)
- := (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~0~1~1~1~1~1~1).
-Notation "'0x7ffe6'" (* 524262 (0x7ffe6) *)
- := (Const 524262%Z).
-Notation "'0x7ffe6'" (* 524262 (0x7ffe6) *)
- := (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~0~0~1~1~0).
-Notation "'0x7ffe7'" (* 524263 (0x7ffe7) *)
- := (Const 524263%Z).
-Notation "'0x7ffe7'" (* 524263 (0x7ffe7) *)
- := (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~0~0~1~1~1).
-Notation "'0x7ffeb'" (* 524267 (0x7ffeb) *)
- := (Const 524267%Z).
-Notation "'0x7ffeb'" (* 524267 (0x7ffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0x7ffed'" (* 524269 (0x7ffed) *)
- := (Const 524269%Z).
-Notation "'0x7ffed'" (* 524269 (0x7ffed) *)
- := (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~0~1~1~0~1).
-Notation "'0x7ffef'" (* 524271 (0x7ffef) *)
- := (Const 524271%Z).
-Notation "'0x7ffef'" (* 524271 (0x7ffef) *)
- := (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~0~1~1~1~1).
-Notation "'0x7fff6'" (* 524278 (0x7fff6) *)
- := (Const 524278%Z).
-Notation "'0x7fff6'" (* 524278 (0x7fff6) *)
- := (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~0~1~1~0).
-Notation "'0x7fff7'" (* 524279 (0x7fff7) *)
- := (Const 524279%Z).
-Notation "'0x7fff7'" (* 524279 (0x7fff7) *)
- := (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~0~1~1~1).
-Notation "'0x7fffe'" (* 524286 (0x7fffe) *)
- := (Const 524286%Z).
-Notation "'0x7fffe'" (* 524286 (0x7fffe) *)
- := (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~0).
-Notation "'0x7ffff'" (* 524287 (0x7ffff) *)
- := (Const 524287%Z).
-Notation "'0x7ffff'" (* 524287 (0x7ffff) *)
- := (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).
-Notation "'0x80000'" (* 524288 (0x80000) *)
- := (Const 524288%Z).
-Notation "'0x80000'" (* 524288 (0x80000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80001'" (* 524289 (0x80001) *)
- := (Const 524289%Z).
-Notation "'0x80001'" (* 524289 (0x80001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xa5fff'" (* 679935 (0xa5fff) *)
- := (Const 679935%Z).
-Notation "'0xa5fff'" (* 679935 (0xa5fff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfc000'" (* 1032192 (0xfc000) *)
- := (Const 1032192%Z).
-Notation "'0xfc000'" (* 1032192 (0xfc000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0xfefff'" (* 1044479 (0xfefff) *)
- := (Const 1044479%Z).
-Notation "'0xfefff'" (* 1044479 (0xfefff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xffc2f'" (* 1047599 (0xffc2f) *)
- := (Const 1047599%Z).
-Notation "'0xffc2f'" (* 1047599 (0xffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0xffd03'" (* 1047811 (0xffd03) *)
- := (Const 1047811%Z).
-Notation "'0xffd03'" (* 1047811 (0xffd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1).
-Notation "'0xffe8a'" (* 1048202 (0xffe8a) *)
- := (Const 1048202%Z).
-Notation "'0xffe8a'" (* 1048202 (0xffe8a) *)
- := (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~0~1~0~0~0~1~0~1~0).
-Notation "'0xfff7e'" (* 1048446 (0xfff7e) *)
- := (Const 1048446%Z).
-Notation "'0xfff7e'" (* 1048446 (0xfff7e) *)
- := (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~0~1~1~1~1~1~1~0).
-Notation "'0xfff97'" (* 1048471 (0xfff97) *)
- := (Const 1048471%Z).
-Notation "'0xfff97'" (* 1048471 (0xfff97) *)
- := (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~0~0~1~0~1~1~1).
-Notation "'0xfffbf'" (* 1048511 (0xfffbf) *)
- := (Const 1048511%Z).
-Notation "'0xfffbf'" (* 1048511 (0xfffbf) *)
- := (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~0~1~1~1~1~1~1).
-Notation "'0xfffce'" (* 1048526 (0xfffce) *)
- := (Const 1048526%Z).
-Notation "'0xfffce'" (* 1048526 (0xfffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0xfffd6'" (* 1048534 (0xfffd6) *)
- := (Const 1048534%Z).
-Notation "'0xfffd6'" (* 1048534 (0xfffd6) *)
- := (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~0~1~0~1~1~0).
-Notation "'0xfffda'" (* 1048538 (0xfffda) *)
- := (Const 1048538%Z).
-Notation "'0xfffda'" (* 1048538 (0xfffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0xfffe5'" (* 1048549 (0xfffe5) *)
- := (Const 1048549%Z).
-Notation "'0xfffe5'" (* 1048549 (0xfffe5) *)
- := (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~0~0~1~0~1).
-Notation "'0xfffed'" (* 1048557 (0xfffed) *)
- := (Const 1048557%Z).
-Notation "'0xfffed'" (* 1048557 (0xfffed) *)
- := (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~0~1~1~0~1).
-Notation "'0xfffee'" (* 1048558 (0xfffee) *)
- := (Const 1048558%Z).
-Notation "'0xfffee'" (* 1048558 (0xfffee) *)
- := (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~0~1~1~1~0).
-Notation "'0xfffef'" (* 1048559 (0xfffef) *)
- := (Const 1048559%Z).
-Notation "'0xfffef'" (* 1048559 (0xfffef) *)
- := (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~0~1~1~1~1).
-Notation "'0xffffd'" (* 1048573 (0xffffd) *)
- := (Const 1048573%Z).
-Notation "'0xffffd'" (* 1048573 (0xffffd) *)
- := (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~0~1).
-Notation "'0xffffe'" (* 1048574 (0xffffe) *)
- := (Const 1048574%Z).
-Notation "'0xffffe'" (* 1048574 (0xffffe) *)
- := (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~0).
-Notation "'0xfffff'" (* 1048575 (0xfffff) *)
- := (Const 1048575%Z).
-Notation "'0xfffff'" (* 1048575 (0xfffff) *)
- := (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).
-Notation "'0x100000'" (* 1048576 (0x100000) *)
- := (Const 1048576%Z).
-Notation "'0x100000'" (* 1048576 (0x100000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100001'" (* 1048577 (0x100001) *)
- := (Const 1048577%Z).
-Notation "'0x100001'" (* 1048577 (0x100001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x14bffe'" (* 1359870 (0x14bffe) *)
- := (Const 1359870%Z).
-Notation "'0x14bffe'" (* 1359870 (0x14bffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x1f6eff'" (* 2060031 (0x1f6eff) *)
- := (Const 2060031%Z).
-Notation "'0x1f6eff'" (* 2060031 (0x1f6eff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~0~1~1~1~1~1~1~1~1).
-Notation "'0x1f8000'" (* 2064384 (0x1f8000) *)
- := (Const 2064384%Z).
-Notation "'0x1f8000'" (* 2064384 (0x1f8000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1fc29f'" (* 2081439 (0x1fc29f) *)
- := (Const 2081439%Z).
-Notation "'0x1fc29f'" (* 2081439 (0x1fc29f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~0~0~1~1~1~1~1).
-Notation "'0x1fdffe'" (* 2088958 (0x1fdffe) *)
- := (Const 2088958%Z).
-Notation "'0x1fdffe'" (* 2088958 (0x1fdffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x1ff4ff'" (* 2094335 (0x1ff4ff) *)
- := (Const 2094335%Z).
-Notation "'0x1ff4ff'" (* 2094335 (0x1ff4ff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~0~0~1~1~1~1~1~1~1~1).
-Notation "'0x1ff85e'" (* 2095198 (0x1ff85e) *)
- := (Const 2095198%Z).
-Notation "'0x1ff85e'" (* 2095198 (0x1ff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0x1ffa06'" (* 2095622 (0x1ffa06) *)
- := (Const 2095622%Z).
-Notation "'0x1ffa06'" (* 2095622 (0x1ffa06) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0x1ffbff'" (* 2096127 (0x1ffbff) *)
- := (Const 2096127%Z).
-Notation "'0x1ffbff'" (* 2096127 (0x1ffbff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x1ffc00'" (* 2096128 (0x1ffc00) *)
- := (Const 2096128%Z).
-Notation "'0x1ffc00'" (* 2096128 (0x1ffc00) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1ffdc7'" (* 2096583 (0x1ffdc7) *)
- := (Const 2096583%Z).
-Notation "'0x1ffdc7'" (* 2096583 (0x1ffdc7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1).
-Notation "'0x1fff2e'" (* 2096942 (0x1fff2e) *)
- := (Const 2096942%Z).
-Notation "'0x1fff2e'" (* 2096942 (0x1fff2e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~1~0).
-Notation "'0x1fffca'" (* 2097098 (0x1fffca) *)
- := (Const 2097098%Z).
-Notation "'0x1fffca'" (* 2097098 (0x1fffca) *)
- := (Const WO~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~0~0~1~0~1~0).
-Notation "'0x1fffda'" (* 2097114 (0x1fffda) *)
- := (Const 2097114%Z).
-Notation "'0x1fffda'" (* 2097114 (0x1fffda) *)
- := (Const WO~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~0~1~1~0~1~0).
-Notation "'0x1fffde'" (* 2097118 (0x1fffde) *)
- := (Const 2097118%Z).
-Notation "'0x1fffde'" (* 2097118 (0x1fffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0x1fffe7'" (* 2097127 (0x1fffe7) *)
- := (Const 2097127%Z).
-Notation "'0x1fffe7'" (* 2097127 (0x1fffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0x1fffef'" (* 2097135 (0x1fffef) *)
- := (Const 2097135%Z).
-Notation "'0x1fffef'" (* 2097135 (0x1fffef) *)
- := (Const WO~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~0~1~1~1~1).
-Notation "'0x1ffff7'" (* 2097143 (0x1ffff7) *)
- := (Const 2097143%Z).
-Notation "'0x1ffff7'" (* 2097143 (0x1ffff7) *)
- := (Const WO~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~0~1~1~1).
-Notation "'0x1ffffe'" (* 2097150 (0x1ffffe) *)
- := (Const 2097150%Z).
-Notation "'0x1ffffe'" (* 2097150 (0x1ffffe) *)
- := (Const WO~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~0).
-Notation "'0x1fffff'" (* 2097151 (0x1fffff) *)
- := (Const 2097151%Z).
-Notation "'0x1fffff'" (* 2097151 (0x1fffff) *)
- := (Const WO~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).
-Notation "'0x200000'" (* 2097152 (0x200000) *)
- := (Const 2097152%Z).
-Notation "'0x200000'" (* 2097152 (0x200000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200001'" (* 2097153 (0x200001) *)
- := (Const 2097153%Z).
-Notation "'0x200001'" (* 2097153 (0x200001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3eddfe'" (* 4120062 (0x3eddfe) *)
- := (Const 4120062%Z).
-Notation "'0x3eddfe'" (* 4120062 (0x3eddfe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~0~1~1~1~1~1~1~1~1~0).
-Notation "'0x3f853e'" (* 4162878 (0x3f853e) *)
- := (Const 4162878%Z).
-Notation "'0x3f853e'" (* 4162878 (0x3f853e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~0~0~1~1~1~1~1~0).
-Notation "'0x3fe9fe'" (* 4188670 (0x3fe9fe) *)
- := (Const 4188670%Z).
-Notation "'0x3fe9fe'" (* 4188670 (0x3fe9fe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~0~0~1~1~1~1~1~1~1~1~0).
-Notation "'0x3ff7fe'" (* 4192254 (0x3ff7fe) *)
- := (Const 4192254%Z).
-Notation "'0x3ff7fe'" (* 4192254 (0x3ff7fe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x3ffb8e'" (* 4193166 (0x3ffb8e) *)
- := (Const 4193166%Z).
-Notation "'0x3ffb8e'" (* 4193166 (0x3ffb8e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0x3ffc2f'" (* 4193327 (0x3ffc2f) *)
- := (Const 4193327%Z).
-Notation "'0x3ffc2f'" (* 4193327 (0x3ffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0x3ffd03'" (* 4193539 (0x3ffd03) *)
- := (Const 4193539%Z).
-Notation "'0x3ffd03'" (* 4193539 (0x3ffd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1).
-Notation "'0x3ffdc7'" (* 4193735 (0x3ffdc7) *)
- := (Const 4193735%Z).
-Notation "'0x3ffdc7'" (* 4193735 (0x3ffdc7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1).
-Notation "'0x3ffe1f'" (* 4193823 (0x3ffe1f) *)
- := (Const 4193823%Z).
-Notation "'0x3ffe1f'" (* 4193823 (0x3ffe1f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1).
-Notation "'0x3ffe5b'" (* 4193883 (0x3ffe5b) *)
- := (Const 4193883%Z).
-Notation "'0x3ffe5b'" (* 4193883 (0x3ffe5b) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1).
-Notation "'0x3ffec3'" (* 4193987 (0x3ffec3) *)
- := (Const 4193987%Z).
-Notation "'0x3ffec3'" (* 4193987 (0x3ffec3) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~0~0~0~1~1).
-Notation "'0x3fff43'" (* 4194115 (0x3fff43) *)
- := (Const 4194115%Z).
-Notation "'0x3fff43'" (* 4194115 (0x3fff43) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~1~1).
-Notation "'0x3fffc0'" (* 4194240 (0x3fffc0) *)
- := (Const 4194240%Z).
-Notation "'0x3fffc0'" (* 4194240 (0x3fffc0) *)
- := (Const WO~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~0~0~0~0~0~0).
-Notation "'0x3fffce'" (* 4194254 (0x3fffce) *)
- := (Const 4194254%Z).
-Notation "'0x3fffce'" (* 4194254 (0x3fffce) *)
- := (Const WO~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~0~0~1~1~1~0).
-Notation "'0x3fffde'" (* 4194270 (0x3fffde) *)
- := (Const 4194270%Z).
-Notation "'0x3fffde'" (* 4194270 (0x3fffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0x3fffdf'" (* 4194271 (0x3fffdf) *)
- := (Const 4194271%Z).
-Notation "'0x3fffdf'" (* 4194271 (0x3fffdf) *)
- := (Const WO~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~0~1~1~1~1~1).
-Notation "'0x3fffe3'" (* 4194275 (0x3fffe3) *)
- := (Const 4194275%Z).
-Notation "'0x3fffe3'" (* 4194275 (0x3fffe3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0x3fffe7'" (* 4194279 (0x3fffe7) *)
- := (Const 4194279%Z).
-Notation "'0x3fffe7'" (* 4194279 (0x3fffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0x3fffed'" (* 4194285 (0x3fffed) *)
- := (Const 4194285%Z).
-Notation "'0x3fffed'" (* 4194285 (0x3fffed) *)
- := (Const WO~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~0~1~1~0~1).
-Notation "'0x3fffee'" (* 4194286 (0x3fffee) *)
- := (Const 4194286%Z).
-Notation "'0x3fffee'" (* 4194286 (0x3fffee) *)
- := (Const WO~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~0~1~1~1~0).
-Notation "'0x3fffef'" (* 4194287 (0x3fffef) *)
- := (Const 4194287%Z).
-Notation "'0x3fffef'" (* 4194287 (0x3fffef) *)
- := (Const WO~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~0~1~1~1~1).
-Notation "'0x3ffff1'" (* 4194289 (0x3ffff1) *)
- := (Const 4194289%Z).
-Notation "'0x3ffff1'" (* 4194289 (0x3ffff1) *)
- := (Const WO~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~0~0~0~1).
-Notation "'0x3ffff5'" (* 4194293 (0x3ffff5) *)
- := (Const 4194293%Z).
-Notation "'0x3ffff5'" (* 4194293 (0x3ffff5) *)
- := (Const WO~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~0~1~0~1).
-Notation "'0x3ffffb'" (* 4194299 (0x3ffffb) *)
- := (Const 4194299%Z).
-Notation "'0x3ffffb'" (* 4194299 (0x3ffffb) *)
- := (Const WO~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~0~1~1).
-Notation "'0x3ffffd'" (* 4194301 (0x3ffffd) *)
- := (Const 4194301%Z).
-Notation "'0x3ffffd'" (* 4194301 (0x3ffffd) *)
- := (Const WO~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~0~1).
-Notation "'0x3ffffe'" (* 4194302 (0x3ffffe) *)
- := (Const 4194302%Z).
-Notation "'0x3ffffe'" (* 4194302 (0x3ffffe) *)
- := (Const WO~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~0).
-Notation "'0x3fffff'" (* 4194303 (0x3fffff) *)
- := (Const 4194303%Z).
-Notation "'0x3fffff'" (* 4194303 (0x3fffff) *)
- := (Const WO~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).
-Notation "'0x400000'" (* 4194304 (0x400000) *)
- := (Const 4194304%Z).
-Notation "'0x400000'" (* 4194304 (0x400000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400001'" (* 4194305 (0x400001) *)
- := (Const 4194305%Z).
-Notation "'0x400001'" (* 4194305 (0x400001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7f01ff'" (* 8323583 (0x7f01ff) *)
- := (Const 8323583%Z).
-Notation "'0x7f01ff'" (* 8323583 (0x7f01ff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1).
-Notation "'0x7fc000'" (* 8372224 (0x7fc000) *)
- := (Const 8372224%Z).
-Notation "'0x7fc000'" (* 8372224 (0x7fc000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x7ff85e'" (* 8386654 (0x7ff85e) *)
- := (Const 8386654%Z).
-Notation "'0x7ff85e'" (* 8386654 (0x7ff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0x7ffa06'" (* 8387078 (0x7ffa06) *)
- := (Const 8387078%Z).
-Notation "'0x7ffa06'" (* 8387078 (0x7ffa06) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0x7ffb8e'" (* 8387470 (0x7ffb8e) *)
- := (Const 8387470%Z).
-Notation "'0x7ffb8e'" (* 8387470 (0x7ffb8e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0x7ffc3e'" (* 8387646 (0x7ffc3e) *)
- := (Const 8387646%Z).
-Notation "'0x7ffc3e'" (* 8387646 (0x7ffc3e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1~0).
-Notation "'0x7ffcb6'" (* 8387766 (0x7ffcb6) *)
- := (Const 8387766%Z).
-Notation "'0x7ffcb6'" (* 8387766 (0x7ffcb6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1~0).
-Notation "'0x7ffd86'" (* 8387974 (0x7ffd86) *)
- := (Const 8387974%Z).
-Notation "'0x7ffd86'" (* 8387974 (0x7ffd86) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~0~0~0~1~1~0).
-Notation "'0x7ffe1f'" (* 8388127 (0x7ffe1f) *)
- := (Const 8388127%Z).
-Notation "'0x7ffe1f'" (* 8388127 (0x7ffe1f) *)
- := (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~0~0~0~0~1~1~1~1~1).
-Notation "'0x7ffe5b'" (* 8388187 (0x7ffe5b) *)
- := (Const 8388187%Z).
-Notation "'0x7ffe5b'" (* 8388187 (0x7ffe5b) *)
- := (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~0~0~1~0~1~1~0~1~1).
-Notation "'0x7ffe86'" (* 8388230 (0x7ffe86) *)
- := (Const 8388230%Z).
-Notation "'0x7ffe86'" (* 8388230 (0x7ffe86) *)
- := (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~0~1~0~0~0~0~1~1~0).
-Notation "'0x7ffec3'" (* 8388291 (0x7ffec3) *)
- := (Const 8388291%Z).
-Notation "'0x7ffec3'" (* 8388291 (0x7ffec3) *)
- := (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~0~1~1~0~0~0~0~1~1).
-Notation "'0x7fff45'" (* 8388421 (0x7fff45) *)
- := (Const 8388421%Z).
-Notation "'0x7fff45'" (* 8388421 (0x7fff45) *)
- := (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~0~1~0~0~0~1~0~1).
-Notation "'0x7fff8b'" (* 8388491 (0x7fff8b) *)
- := (Const 8388491%Z).
-Notation "'0x7fff8b'" (* 8388491 (0x7fff8b) *)
- := (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~0~0~0~1~0~1~1).
-Notation "'0x7fff97'" (* 8388503 (0x7fff97) *)
- := (Const 8388503%Z).
-Notation "'0x7fff97'" (* 8388503 (0x7fff97) *)
- := (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~0~0~1~0~1~1~1).
-Notation "'0x7fffbe'" (* 8388542 (0x7fffbe) *)
- := (Const 8388542%Z).
-Notation "'0x7fffbe'" (* 8388542 (0x7fffbe) *)
- := (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~0~1~1~1~1~1~0).
-Notation "'0x7fffce'" (* 8388558 (0x7fffce) *)
- := (Const 8388558%Z).
-Notation "'0x7fffce'" (* 8388558 (0x7fffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0x7fffda'" (* 8388570 (0x7fffda) *)
- := (Const 8388570%Z).
-Notation "'0x7fffda'" (* 8388570 (0x7fffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0x7fffde'" (* 8388574 (0x7fffde) *)
- := (Const 8388574%Z).
-Notation "'0x7fffde'" (* 8388574 (0x7fffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0x7fffe1'" (* 8388577 (0x7fffe1) *)
- := (Const 8388577%Z).
-Notation "'0x7fffe1'" (* 8388577 (0x7fffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0x7fffe2'" (* 8388578 (0x7fffe2) *)
- := (Const 8388578%Z).
-Notation "'0x7fffe2'" (* 8388578 (0x7fffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0x7fffe5'" (* 8388581 (0x7fffe5) *)
- := (Const 8388581%Z).
-Notation "'0x7fffe5'" (* 8388581 (0x7fffe5) *)
- := (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~0~0~1~0~1).
-Notation "'0x7fffef'" (* 8388591 (0x7fffef) *)
- := (Const 8388591%Z).
-Notation "'0x7fffef'" (* 8388591 (0x7fffef) *)
- := (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~0~1~1~1~1).
-Notation "'0x7ffff1'" (* 8388593 (0x7ffff1) *)
- := (Const 8388593%Z).
-Notation "'0x7ffff1'" (* 8388593 (0x7ffff1) *)
- := (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~0~0~0~1).
-Notation "'0x7ffff3'" (* 8388595 (0x7ffff3) *)
- := (Const 8388595%Z).
-Notation "'0x7ffff3'" (* 8388595 (0x7ffff3) *)
- := (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~0~0~1~1).
-Notation "'0x7ffff6'" (* 8388598 (0x7ffff6) *)
- := (Const 8388598%Z).
-Notation "'0x7ffff6'" (* 8388598 (0x7ffff6) *)
- := (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~0~1~1~0).
-Notation "'0x7ffff7'" (* 8388599 (0x7ffff7) *)
- := (Const 8388599%Z).
-Notation "'0x7ffff7'" (* 8388599 (0x7ffff7) *)
- := (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~0~1~1~1).
-Notation "'0x7ffffa'" (* 8388602 (0x7ffffa) *)
- := (Const 8388602%Z).
-Notation "'0x7ffffa'" (* 8388602 (0x7ffffa) *)
- := (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~0~1~0).
-Notation "'0x7ffffb'" (* 8388603 (0x7ffffb) *)
- := (Const 8388603%Z).
-Notation "'0x7ffffb'" (* 8388603 (0x7ffffb) *)
- := (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~0~1~1).
-Notation "'0x7ffffc'" (* 8388604 (0x7ffffc) *)
- := (Const 8388604%Z).
-Notation "'0x7ffffc'" (* 8388604 (0x7ffffc) *)
- := (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~0~0).
-Notation "'0x7ffffd'" (* 8388605 (0x7ffffd) *)
- := (Const 8388605%Z).
-Notation "'0x7ffffd'" (* 8388605 (0x7ffffd) *)
- := (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~0~1).
-Notation "'0x7ffffe'" (* 8388606 (0x7ffffe) *)
- := (Const 8388606%Z).
-Notation "'0x7ffffe'" (* 8388606 (0x7ffffe) *)
- := (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~0).
-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 "'0x800000'" (* 8388608 (0x800000) *)
- := (Const 8388608%Z).
-Notation "'0x800000'" (* 8388608 (0x800000) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800001'" (* 8388609 (0x800001) *)
- := (Const 8388609%Z).
-Notation "'0x800001'" (* 8388609 (0x800001) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xb0ffff'" (* 11599871 (0xb0ffff) *)
- := (Const 11599871%Z).
-Notation "'0xb0ffff'" (* 11599871 (0xb0ffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~1~1~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfe03fe'" (* 16647166 (0xfe03fe) *)
- := (Const 16647166%Z).
-Notation "'0xfe03fe'" (* 16647166 (0xfe03fe) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0).
-Notation "'0xfeffff'" (* 16711679 (0xfeffff) *)
- := (Const 16711679%Z).
-Notation "'0xfeffff'" (* 16711679 (0xfeffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfffaff'" (* 16775935 (0xfffaff) *)
- := (Const 16775935%Z).
-Notation "'0xfffaff'" (* 16775935 (0xfffaff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~1~1~1~1~1~1~1~1).
-Notation "'0xfffc3e'" (* 16776254 (0xfffc3e) *)
- := (Const 16776254%Z).
-Notation "'0xfffc3e'" (* 16776254 (0xfffc3e) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1~0).
-Notation "'0xfffcb6'" (* 16776374 (0xfffcb6) *)
- := (Const 16776374%Z).
-Notation "'0xfffcb6'" (* 16776374 (0xfffcb6) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1~0).
-Notation "'0xfffd86'" (* 16776582 (0xfffd86) *)
- := (Const 16776582%Z).
-Notation "'0xfffd86'" (* 16776582 (0xfffd86) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~0~0~0~0~1~1~0).
-Notation "'0xfffe8a'" (* 16776842 (0xfffe8a) *)
- := (Const 16776842%Z).
-Notation "'0xfffe8a'" (* 16776842 (0xfffe8a) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~1~0~1~0).
-Notation "'0xfffeff'" (* 16776959 (0xfffeff) *)
- := (Const 16776959%Z).
-Notation "'0xfffeff'" (* 16776959 (0xfffeff) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1).
-Notation "'0xffff16'" (* 16776982 (0xffff16) *)
- := (Const 16776982%Z).
-Notation "'0xffff16'" (* 16776982 (0xffff16) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~1~0~1~1~0).
-Notation "'0xffff2e'" (* 16777006 (0xffff2e) *)
- := (Const 16777006%Z).
-Notation "'0xffff2e'" (* 16777006 (0xffff2e) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~1~0).
-Notation "'0xffff43'" (* 16777027 (0xffff43) *)
- := (Const 16777027%Z).
-Notation "'0xffff43'" (* 16777027 (0xffff43) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~0~1~1).
-Notation "'0xffff45'" (* 16777029 (0xffff45) *)
- := (Const 16777029%Z).
-Notation "'0xffff45'" (* 16777029 (0xffff45) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~0~0~1~0~1).
-Notation "'0xffff97'" (* 16777111 (0xffff97) *)
- := (Const 16777111%Z).
-Notation "'0xffff97'" (* 16777111 (0xffff97) *)
- := (Const WO~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~0~0~1~0~1~1~1).
-Notation "'0xffffca'" (* 16777162 (0xffffca) *)
- := (Const 16777162%Z).
-Notation "'0xffffca'" (* 16777162 (0xffffca) *)
- := (Const WO~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~0~0~1~0~1~0).
-Notation "'0xffffde'" (* 16777182 (0xffffde) *)
- := (Const 16777182%Z).
-Notation "'0xffffde'" (* 16777182 (0xffffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0xffffe1'" (* 16777185 (0xffffe1) *)
- := (Const 16777185%Z).
-Notation "'0xffffe1'" (* 16777185 (0xffffe1) *)
- := (Const WO~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~0~0~0~0~1).
-Notation "'0xffffe2'" (* 16777186 (0xffffe2) *)
- := (Const 16777186%Z).
-Notation "'0xffffe2'" (* 16777186 (0xffffe2) *)
- := (Const WO~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~0~0~0~1~0).
-Notation "'0xffffe3'" (* 16777187 (0xffffe3) *)
- := (Const 16777187%Z).
-Notation "'0xffffe3'" (* 16777187 (0xffffe3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0xffffe5'" (* 16777189 (0xffffe5) *)
- := (Const 16777189%Z).
-Notation "'0xffffe5'" (* 16777189 (0xffffe5) *)
- := (Const WO~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~0~0~1~0~1).
-Notation "'0xffffe6'" (* 16777190 (0xffffe6) *)
- := (Const 16777190%Z).
-Notation "'0xffffe6'" (* 16777190 (0xffffe6) *)
- := (Const WO~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~0~0~1~1~0).
-Notation "'0xffffe7'" (* 16777191 (0xffffe7) *)
- := (Const 16777191%Z).
-Notation "'0xffffe7'" (* 16777191 (0xffffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0xffffed'" (* 16777197 (0xffffed) *)
- := (Const 16777197%Z).
-Notation "'0xffffed'" (* 16777197 (0xffffed) *)
- := (Const WO~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~0~1~1~0~1).
-Notation "'0xffffee'" (* 16777198 (0xffffee) *)
- := (Const 16777198%Z).
-Notation "'0xffffee'" (* 16777198 (0xffffee) *)
- := (Const WO~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~0~1~1~1~0).
-Notation "'0xffffef'" (* 16777199 (0xffffef) *)
- := (Const 16777199%Z).
-Notation "'0xffffef'" (* 16777199 (0xffffef) *)
- := (Const WO~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~0~1~1~1~1).
-Notation "'0xfffff1'" (* 16777201 (0xfffff1) *)
- := (Const 16777201%Z).
-Notation "'0xfffff1'" (* 16777201 (0xfffff1) *)
- := (Const WO~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~0~0~0~1).
-Notation "'0xfffff5'" (* 16777205 (0xfffff5) *)
- := (Const 16777205%Z).
-Notation "'0xfffff5'" (* 16777205 (0xfffff5) *)
- := (Const WO~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~0~1~0~1).
-Notation "'0xfffff6'" (* 16777206 (0xfffff6) *)
- := (Const 16777206%Z).
-Notation "'0xfffff6'" (* 16777206 (0xfffff6) *)
- := (Const WO~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~0~1~1~0).
-Notation "'0xfffff7'" (* 16777207 (0xfffff7) *)
- := (Const 16777207%Z).
-Notation "'0xfffff7'" (* 16777207 (0xfffff7) *)
- := (Const WO~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~0~1~1~1).
-Notation "'0xfffffa'" (* 16777210 (0xfffffa) *)
- := (Const 16777210%Z).
-Notation "'0xfffffa'" (* 16777210 (0xfffffa) *)
- := (Const WO~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~0~1~0).
-Notation "'0xfffffb'" (* 16777211 (0xfffffb) *)
- := (Const 16777211%Z).
-Notation "'0xfffffb'" (* 16777211 (0xfffffb) *)
- := (Const WO~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~0~1~1).
-Notation "'0xfffffc'" (* 16777212 (0xfffffc) *)
- := (Const 16777212%Z).
-Notation "'0xfffffc'" (* 16777212 (0xfffffc) *)
- := (Const WO~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~0~0).
-Notation "'0xfffffd'" (* 16777213 (0xfffffd) *)
- := (Const 16777213%Z).
-Notation "'0xfffffd'" (* 16777213 (0xfffffd) *)
- := (Const WO~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~0~1).
-Notation "'0xfffffe'" (* 16777214 (0xfffffe) *)
- := (Const 16777214%Z).
-Notation "'0xfffffe'" (* 16777214 (0xfffffe) *)
- := (Const WO~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~0).
-Notation "'0xffffff'" (* 16777215 (0xffffff) *)
- := (Const 16777215%Z).
-Notation "'0xffffff'" (* 16777215 (0xffffff) *)
- := (Const WO~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).
-Notation "'0x1000000'" (* 16777216 (0x1000000) *)
- := (Const 16777216%Z).
-Notation "'0x1000000'" (* 16777216 (0x1000000) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000001'" (* 16777217 (0x1000001) *)
- := (Const 16777217%Z).
-Notation "'0x1000001'" (* 16777217 (0x1000001) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x161fffe'" (* 23199742 (0x161fffe) *)
- := (Const 23199742%Z).
-Notation "'0x161fffe'" (* 23199742 (0x161fffe) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~1~1~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x1fdfffe'" (* 33423358 (0x1fdfffe) *)
- := (Const 33423358%Z).
-Notation "'0x1fdfffe'" (* 33423358 (0x1fdfffe) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x1fff5fe'" (* 33551870 (0x1fff5fe) *)
- := (Const 33551870%Z).
-Notation "'0x1fff5fe'" (* 33551870 (0x1fff5fe) *)
- := (Const WO~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~1~1~1~1~1~1~1~1~0).
-Notation "'0x1fffe86'" (* 33554054 (0x1fffe86) *)
- := (Const 33554054%Z).
-Notation "'0x1fffe86'" (* 33554054 (0x1fffe86) *)
- := (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~0~1~0~0~0~0~1~1~0).
-Notation "'0x1fffe8a'" (* 33554058 (0x1fffe8a) *)
- := (Const 33554058%Z).
-Notation "'0x1fffe8a'" (* 33554058 (0x1fffe8a) *)
- := (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~0~1~0~0~0~1~0~1~0).
-Notation "'0x1ffff2e'" (* 33554222 (0x1ffff2e) *)
- := (Const 33554222%Z).
-Notation "'0x1ffff2e'" (* 33554222 (0x1ffff2e) *)
- := (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~0~0~1~0~1~1~1~0).
-Notation "'0x1ffff8b'" (* 33554315 (0x1ffff8b) *)
- := (Const 33554315%Z).
-Notation "'0x1ffff8b'" (* 33554315 (0x1ffff8b) *)
- := (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~0~0~0~1~0~1~1).
-Notation "'0x1ffffc2'" (* 33554370 (0x1ffffc2) *)
- := (Const 33554370%Z).
-Notation "'0x1ffffc2'" (* 33554370 (0x1ffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0x1ffffc6'" (* 33554374 (0x1ffffc6) *)
- := (Const 33554374%Z).
-Notation "'0x1ffffc6'" (* 33554374 (0x1ffffc6) *)
- := (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~0~0~0~1~1~0).
-Notation "'0x1ffffca'" (* 33554378 (0x1ffffca) *)
- := (Const 33554378%Z).
-Notation "'0x1ffffca'" (* 33554378 (0x1ffffca) *)
- := (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~0~0~1~0~1~0).
-Notation "'0x1ffffce'" (* 33554382 (0x1ffffce) *)
- := (Const 33554382%Z).
-Notation "'0x1ffffce'" (* 33554382 (0x1ffffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0x1ffffda'" (* 33554394 (0x1ffffda) *)
- := (Const 33554394%Z).
-Notation "'0x1ffffda'" (* 33554394 (0x1ffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0x1ffffde'" (* 33554398 (0x1ffffde) *)
- := (Const 33554398%Z).
-Notation "'0x1ffffde'" (* 33554398 (0x1ffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0x1ffffdf'" (* 33554399 (0x1ffffdf) *)
- := (Const 33554399%Z).
-Notation "'0x1ffffdf'" (* 33554399 (0x1ffffdf) *)
- := (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~0~1~1~1~1~1).
-Notation "'0x1ffffe1'" (* 33554401 (0x1ffffe1) *)
- := (Const 33554401%Z).
-Notation "'0x1ffffe1'" (* 33554401 (0x1ffffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0x1ffffe2'" (* 33554402 (0x1ffffe2) *)
- := (Const 33554402%Z).
-Notation "'0x1ffffe2'" (* 33554402 (0x1ffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0x1ffffe7'" (* 33554407 (0x1ffffe7) *)
- := (Const 33554407%Z).
-Notation "'0x1ffffe7'" (* 33554407 (0x1ffffe7) *)
- := (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~0~0~1~1~1).
-Notation "'0x1ffffea'" (* 33554410 (0x1ffffea) *)
- := (Const 33554410%Z).
-Notation "'0x1ffffea'" (* 33554410 (0x1ffffea) *)
- := (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~0~1~0~1~0).
-Notation "'0x1ffffeb'" (* 33554411 (0x1ffffeb) *)
- := (Const 33554411%Z).
-Notation "'0x1ffffeb'" (* 33554411 (0x1ffffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0x1ffffed'" (* 33554413 (0x1ffffed) *)
- := (Const 33554413%Z).
-Notation "'0x1ffffed'" (* 33554413 (0x1ffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0x1ffffee'" (* 33554414 (0x1ffffee) *)
- := (Const 33554414%Z).
-Notation "'0x1ffffee'" (* 33554414 (0x1ffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0x1ffffef'" (* 33554415 (0x1ffffef) *)
- := (Const 33554415%Z).
-Notation "'0x1ffffef'" (* 33554415 (0x1ffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0x1fffff1'" (* 33554417 (0x1fffff1) *)
- := (Const 33554417%Z).
-Notation "'0x1fffff1'" (* 33554417 (0x1fffff1) *)
- := (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~0~0~0~1).
-Notation "'0x1fffff6'" (* 33554422 (0x1fffff6) *)
- := (Const 33554422%Z).
-Notation "'0x1fffff6'" (* 33554422 (0x1fffff6) *)
- := (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~0~1~1~0).
-Notation "'0x1fffff7'" (* 33554423 (0x1fffff7) *)
- := (Const 33554423%Z).
-Notation "'0x1fffff7'" (* 33554423 (0x1fffff7) *)
- := (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~0~1~1~1).
-Notation "'0x1fffffa'" (* 33554426 (0x1fffffa) *)
- := (Const 33554426%Z).
-Notation "'0x1fffffa'" (* 33554426 (0x1fffffa) *)
- := (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~0~1~0).
-Notation "'0x1fffffb'" (* 33554427 (0x1fffffb) *)
- := (Const 33554427%Z).
-Notation "'0x1fffffb'" (* 33554427 (0x1fffffb) *)
- := (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~0~1~1).
-Notation "'0x1fffffc'" (* 33554428 (0x1fffffc) *)
- := (Const 33554428%Z).
-Notation "'0x1fffffc'" (* 33554428 (0x1fffffc) *)
- := (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~0~0).
-Notation "'0x1fffffd'" (* 33554429 (0x1fffffd) *)
- := (Const 33554429%Z).
-Notation "'0x1fffffd'" (* 33554429 (0x1fffffd) *)
- := (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~0~1).
-Notation "'0x1fffffe'" (* 33554430 (0x1fffffe) *)
- := (Const 33554430%Z).
-Notation "'0x1fffffe'" (* 33554430 (0x1fffffe) *)
- := (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~0).
-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 "'0x2000000'" (* 33554432 (0x2000000) *)
- := (Const 33554432%Z).
-Notation "'0x2000000'" (* 33554432 (0x2000000) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000001'" (* 33554433 (0x2000001) *)
- := (Const 33554433%Z).
-Notation "'0x2000001'" (* 33554433 (0x2000001) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffc2f'" (* 67107887 (0x3fffc2f) *)
- := (Const 67107887%Z).
-Notation "'0x3fffc2f'" (* 67107887 (0x3fffc2f) *)
- := (Const WO~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0x3ffff16'" (* 67108630 (0x3ffff16) *)
- := (Const 67108630%Z).
-Notation "'0x3ffff16'" (* 67108630 (0x3ffff16) *)
- := (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~0~0~0~1~0~1~1~0).
-Notation "'0x3ffffbe'" (* 67108798 (0x3ffffbe) *)
- := (Const 67108798%Z).
-Notation "'0x3ffffbe'" (* 67108798 (0x3ffffbe) *)
- := (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~0~1~1~1~1~1~0).
-Notation "'0x3ffffbf'" (* 67108799 (0x3ffffbf) *)
- := (Const 67108799%Z).
-Notation "'0x3ffffbf'" (* 67108799 (0x3ffffbf) *)
- := (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~0~1~1~1~1~1~1).
-Notation "'0x3ffffc2'" (* 67108802 (0x3ffffc2) *)
- := (Const 67108802%Z).
-Notation "'0x3ffffc2'" (* 67108802 (0x3ffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0x3ffffce'" (* 67108814 (0x3ffffce) *)
- := (Const 67108814%Z).
-Notation "'0x3ffffce'" (* 67108814 (0x3ffffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0x3ffffd6'" (* 67108822 (0x3ffffd6) *)
- := (Const 67108822%Z).
-Notation "'0x3ffffd6'" (* 67108822 (0x3ffffd6) *)
- := (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~0~1~0~1~1~0).
-Notation "'0x3ffffda'" (* 67108826 (0x3ffffda) *)
- := (Const 67108826%Z).
-Notation "'0x3ffffda'" (* 67108826 (0x3ffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0x3ffffde'" (* 67108830 (0x3ffffde) *)
- := (Const 67108830%Z).
-Notation "'0x3ffffde'" (* 67108830 (0x3ffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0x3ffffe1'" (* 67108833 (0x3ffffe1) *)
- := (Const 67108833%Z).
-Notation "'0x3ffffe1'" (* 67108833 (0x3ffffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0x3ffffe2'" (* 67108834 (0x3ffffe2) *)
- := (Const 67108834%Z).
-Notation "'0x3ffffe2'" (* 67108834 (0x3ffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0x3ffffe5'" (* 67108837 (0x3ffffe5) *)
- := (Const 67108837%Z).
-Notation "'0x3ffffe5'" (* 67108837 (0x3ffffe5) *)
- := (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~0~0~1~0~1).
-Notation "'0x3ffffe7'" (* 67108839 (0x3ffffe7) *)
- := (Const 67108839%Z).
-Notation "'0x3ffffe7'" (* 67108839 (0x3ffffe7) *)
- := (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~0~0~1~1~1).
-Notation "'0x3ffffeb'" (* 67108843 (0x3ffffeb) *)
- := (Const 67108843%Z).
-Notation "'0x3ffffeb'" (* 67108843 (0x3ffffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0x3ffffed'" (* 67108845 (0x3ffffed) *)
- := (Const 67108845%Z).
-Notation "'0x3ffffed'" (* 67108845 (0x3ffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0x3ffffee'" (* 67108846 (0x3ffffee) *)
- := (Const 67108846%Z).
-Notation "'0x3ffffee'" (* 67108846 (0x3ffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0x3ffffef'" (* 67108847 (0x3ffffef) *)
- := (Const 67108847%Z).
-Notation "'0x3ffffef'" (* 67108847 (0x3ffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0x3fffff1'" (* 67108849 (0x3fffff1) *)
- := (Const 67108849%Z).
-Notation "'0x3fffff1'" (* 67108849 (0x3fffff1) *)
- := (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~0~0~0~1).
-Notation "'0x3fffff6'" (* 67108854 (0x3fffff6) *)
- := (Const 67108854%Z).
-Notation "'0x3fffff6'" (* 67108854 (0x3fffff6) *)
- := (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~0~1~1~0).
-Notation "'0x3fffff7'" (* 67108855 (0x3fffff7) *)
- := (Const 67108855%Z).
-Notation "'0x3fffff7'" (* 67108855 (0x3fffff7) *)
- := (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~0~1~1~1).
-Notation "'0x3fffffa'" (* 67108858 (0x3fffffa) *)
- := (Const 67108858%Z).
-Notation "'0x3fffffa'" (* 67108858 (0x3fffffa) *)
- := (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~0~1~0).
-Notation "'0x3fffffb'" (* 67108859 (0x3fffffb) *)
- := (Const 67108859%Z).
-Notation "'0x3fffffb'" (* 67108859 (0x3fffffb) *)
- := (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~0~1~1).
-Notation "'0x3fffffc'" (* 67108860 (0x3fffffc) *)
- := (Const 67108860%Z).
-Notation "'0x3fffffc'" (* 67108860 (0x3fffffc) *)
- := (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~0~0).
-Notation "'0x3fffffd'" (* 67108861 (0x3fffffd) *)
- := (Const 67108861%Z).
-Notation "'0x3fffffd'" (* 67108861 (0x3fffffd) *)
- := (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~0~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 "'0x4000000'" (* 67108864 (0x4000000) *)
- := (Const 67108864%Z).
-Notation "'0x4000000'" (* 67108864 (0x4000000) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000001'" (* 67108865 (0x4000001) *)
- := (Const 67108865%Z).
-Notation "'0x4000001'" (* 67108865 (0x4000001) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fff85e'" (* 134215774 (0x7fff85e) *)
- := (Const 134215774%Z).
-Notation "'0x7fff85e'" (* 134215774 (0x7fff85e) *)
- := (Const WO~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0x7ffff7e'" (* 134217598 (0x7ffff7e) *)
- := (Const 134217598%Z).
-Notation "'0x7ffff7e'" (* 134217598 (0x7ffff7e) *)
- := (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~0~1~1~1~1~1~1~0).
-Notation "'0x7ffffc2'" (* 134217666 (0x7ffffc2) *)
- := (Const 134217666%Z).
-Notation "'0x7ffffc2'" (* 134217666 (0x7ffffc2) *)
- := (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~0~0~0~1~0).
-Notation "'0x7ffffca'" (* 134217674 (0x7ffffca) *)
- := (Const 134217674%Z).
-Notation "'0x7ffffca'" (* 134217674 (0x7ffffca) *)
- := (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~0~1~0~1~0).
-Notation "'0x7ffffce'" (* 134217678 (0x7ffffce) *)
- := (Const 134217678%Z).
-Notation "'0x7ffffce'" (* 134217678 (0x7ffffce) *)
- := (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~0~1~1~1~0).
-Notation "'0x7ffffd6'" (* 134217686 (0x7ffffd6) *)
- := (Const 134217686%Z).
-Notation "'0x7ffffd6'" (* 134217686 (0x7ffffd6) *)
- := (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~0~1~1~0).
-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 "'0x7ffffde'" (* 134217694 (0x7ffffde) *)
- := (Const 134217694%Z).
-Notation "'0x7ffffde'" (* 134217694 (0x7ffffde) *)
- := (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~1~1~0).
-Notation "'0x7ffffe1'" (* 134217697 (0x7ffffe1) *)
- := (Const 134217697%Z).
-Notation "'0x7ffffe1'" (* 134217697 (0x7ffffe1) *)
- := (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~0~0~0~1).
-Notation "'0x7ffffe2'" (* 134217698 (0x7ffffe2) *)
- := (Const 134217698%Z).
-Notation "'0x7ffffe2'" (* 134217698 (0x7ffffe2) *)
- := (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~0~0~1~0).
-Notation "'0x7ffffe3'" (* 134217699 (0x7ffffe3) *)
- := (Const 134217699%Z).
-Notation "'0x7ffffe3'" (* 134217699 (0x7ffffe3) *)
- := (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~0~0~1~1).
-Notation "'0x7ffffe7'" (* 134217703 (0x7ffffe7) *)
- := (Const 134217703%Z).
-Notation "'0x7ffffe7'" (* 134217703 (0x7ffffe7) *)
- := (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~0~1~1~1).
-Notation "'0x7ffffed'" (* 134217709 (0x7ffffed) *)
- := (Const 134217709%Z).
-Notation "'0x7ffffed'" (* 134217709 (0x7ffffed) *)
- := (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~0~1).
-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 "'0x7ffffef'" (* 134217711 (0x7ffffef) *)
- := (Const 134217711%Z).
-Notation "'0x7ffffef'" (* 134217711 (0x7ffffef) *)
- := (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~1).
-Notation "'0x7fffff1'" (* 134217713 (0x7fffff1) *)
- := (Const 134217713%Z).
-Notation "'0x7fffff1'" (* 134217713 (0x7fffff1) *)
- := (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~0~0~0~1).
-Notation "'0x7fffff6'" (* 134217718 (0x7fffff6) *)
- := (Const 134217718%Z).
-Notation "'0x7fffff6'" (* 134217718 (0x7fffff6) *)
- := (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~0~1~1~0).
-Notation "'0x7fffff7'" (* 134217719 (0x7fffff7) *)
- := (Const 134217719%Z).
-Notation "'0x7fffff7'" (* 134217719 (0x7fffff7) *)
- := (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~0~1~1~1).
-Notation "'0x7fffffa'" (* 134217722 (0x7fffffa) *)
- := (Const 134217722%Z).
-Notation "'0x7fffffa'" (* 134217722 (0x7fffffa) *)
- := (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~0~1~0).
-Notation "'0x7fffffc'" (* 134217724 (0x7fffffc) *)
- := (Const 134217724%Z).
-Notation "'0x7fffffc'" (* 134217724 (0x7fffffc) *)
- := (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~0~0).
-Notation "'0x7fffffd'" (* 134217725 (0x7fffffd) *)
- := (Const 134217725%Z).
-Notation "'0x7fffffd'" (* 134217725 (0x7fffffd) *)
- := (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~0~1).
-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 "'0x8000000'" (* 134217728 (0x8000000) *)
- := (Const 134217728%Z).
-Notation "'0x8000000'" (* 134217728 (0x8000000) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000001'" (* 134217729 (0x8000001) *)
- := (Const 134217729%Z).
-Notation "'0x8000001'" (* 134217729 (0x8000001) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffff000'" (* 268431360 (0xffff000) *)
- := (Const 268431360%Z).
-Notation "'0xffff000'" (* 268431360 (0xffff000) *)
- := (Const WO~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0xfffffc2'" (* 268435394 (0xfffffc2) *)
- := (Const 268435394%Z).
-Notation "'0xfffffc2'" (* 268435394 (0xfffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0xfffffc6'" (* 268435398 (0xfffffc6) *)
- := (Const 268435398%Z).
-Notation "'0xfffffc6'" (* 268435398 (0xfffffc6) *)
- := (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~0~0~0~1~1~0).
-Notation "'0xfffffce'" (* 268435406 (0xfffffce) *)
- := (Const 268435406%Z).
-Notation "'0xfffffce'" (* 268435406 (0xfffffce) *)
- := (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~0~0~1~1~1~0).
-Notation "'0xfffffda'" (* 268435418 (0xfffffda) *)
- := (Const 268435418%Z).
-Notation "'0xfffffda'" (* 268435418 (0xfffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0xfffffde'" (* 268435422 (0xfffffde) *)
- := (Const 268435422%Z).
-Notation "'0xfffffde'" (* 268435422 (0xfffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0xfffffe2'" (* 268435426 (0xfffffe2) *)
- := (Const 268435426%Z).
-Notation "'0xfffffe2'" (* 268435426 (0xfffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0xfffffee'" (* 268435438 (0xfffffee) *)
- := (Const 268435438%Z).
-Notation "'0xfffffee'" (* 268435438 (0xfffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0xffffff1'" (* 268435441 (0xffffff1) *)
- := (Const 268435441%Z).
-Notation "'0xffffff1'" (* 268435441 (0xffffff1) *)
- := (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~0~0~0~1).
-Notation "'0xffffff3'" (* 268435443 (0xffffff3) *)
- := (Const 268435443%Z).
-Notation "'0xffffff3'" (* 268435443 (0xffffff3) *)
- := (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~0~0~1~1).
-Notation "'0xffffff5'" (* 268435445 (0xffffff5) *)
- := (Const 268435445%Z).
-Notation "'0xffffff5'" (* 268435445 (0xffffff5) *)
- := (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~0~1~0~1).
-Notation "'0xffffff7'" (* 268435447 (0xffffff7) *)
- := (Const 268435447%Z).
-Notation "'0xffffff7'" (* 268435447 (0xffffff7) *)
- := (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~0~1~1~1).
-Notation "'0xffffffa'" (* 268435450 (0xffffffa) *)
- := (Const 268435450%Z).
-Notation "'0xffffffa'" (* 268435450 (0xffffffa) *)
- := (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~0~1~0).
-Notation "'0xffffffb'" (* 268435451 (0xffffffb) *)
- := (Const 268435451%Z).
-Notation "'0xffffffb'" (* 268435451 (0xffffffb) *)
- := (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~0~1~1).
-Notation "'0xffffffc'" (* 268435452 (0xffffffc) *)
- := (Const 268435452%Z).
-Notation "'0xffffffc'" (* 268435452 (0xffffffc) *)
- := (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~0~0).
-Notation "'0xffffffd'" (* 268435453 (0xffffffd) *)
- := (Const 268435453%Z).
-Notation "'0xffffffd'" (* 268435453 (0xffffffd) *)
- := (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~0~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 "'0x10000000'" (* 268435456 (0x10000000) *)
- := (Const 268435456%Z).
-Notation "'0x10000000'" (* 268435456 (0x10000000) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000001'" (* 268435457 (0x10000001) *)
- := (Const 268435457%Z).
-Notation "'0x10000001'" (* 268435457 (0x10000001) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffc2f'" (* 536869935 (0x1ffffc2f) *)
- := (Const 536869935%Z).
-Notation "'0x1ffffc2f'" (* 536869935 (0x1ffffc2f) *)
- := (Const WO~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0x1fffffe2'" (* 536870882 (0x1fffffe2) *)
- := (Const 536870882%Z).
-Notation "'0x1fffffe2'" (* 536870882 (0x1fffffe2) *)
- := (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~0~0~0~1~0).
-Notation "'0x1fffffe6'" (* 536870886 (0x1fffffe6) *)
- := (Const 536870886%Z).
-Notation "'0x1fffffe6'" (* 536870886 (0x1fffffe6) *)
- := (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~0~0~1~1~0).
-Notation "'0x1fffffea'" (* 536870890 (0x1fffffea) *)
- := (Const 536870890%Z).
-Notation "'0x1fffffea'" (* 536870890 (0x1fffffea) *)
- := (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~0~1~0~1~0).
-Notation "'0x1fffffed'" (* 536870893 (0x1fffffed) *)
- := (Const 536870893%Z).
-Notation "'0x1fffffed'" (* 536870893 (0x1fffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0x1fffffee'" (* 536870894 (0x1fffffee) *)
- := (Const 536870894%Z).
-Notation "'0x1fffffee'" (* 536870894 (0x1fffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0x1ffffff6'" (* 536870902 (0x1ffffff6) *)
- := (Const 536870902%Z).
-Notation "'0x1ffffff6'" (* 536870902 (0x1ffffff6) *)
- := (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~0~1~1~0).
-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 "'0x1ffffffb'" (* 536870907 (0x1ffffffb) *)
- := (Const 536870907%Z).
-Notation "'0x1ffffffb'" (* 536870907 (0x1ffffffb) *)
- := (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~1).
-Notation "'0x1ffffffc'" (* 536870908 (0x1ffffffc) *)
- := (Const 536870908%Z).
-Notation "'0x1ffffffc'" (* 536870908 (0x1ffffffc) *)
- := (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~0~0).
-Notation "'0x1ffffffd'" (* 536870909 (0x1ffffffd) *)
- := (Const 536870909%Z).
-Notation "'0x1ffffffd'" (* 536870909 (0x1ffffffd) *)
- := (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~0~1).
-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 "'0x1fffffff'" (* 536870911 (0x1fffffff) *)
- := (Const 536870911%Z).
-Notation "'0x1fffffff'" (* 536870911 (0x1fffffff) *)
- := (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~1).
-Notation "'0x20000000'" (* 536870912 (0x20000000) *)
- := (Const 536870912%Z).
-Notation "'0x20000000'" (* 536870912 (0x20000000) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000001'" (* 536870913 (0x20000001) *)
- := (Const 536870913%Z).
-Notation "'0x20000001'" (* 536870913 (0x20000001) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x286bca1b'" (* 678152731 (0x286bca1b) *)
- := (Const 678152731%Z).
-Notation "'0x286bca1b'" (* 678152731 (0x286bca1b) *)
- := (Const WO~0~0~1~0~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~1).
-Notation "'0x38e38e39'" (* 954437177 (0x38e38e39) *)
- := (Const 954437177%Z).
-Notation "'0x38e38e39'" (* 954437177 (0x38e38e39) *)
- := (Const WO~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~1).
-Notation "'0x3eddffff'" (* 1054736383 (0x3eddffff) *)
- := (Const 1054736383%Z).
-Notation "'0x3eddffff'" (* 1054736383 (0x3eddffff) *)
- := (Const WO~0~0~1~1~1~1~1~0~1~1~0~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x3f80ffff'" (* 1065418751 (0x3f80ffff) *)
- := (Const 1065418751%Z).
-Notation "'0x3f80ffff'" (* 1065418751 (0x3f80ffff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x3fff7fff'" (* 1073709055 (0x3fff7fff) *)
- := (Const 1073709055%Z).
-Notation "'0x3fff7fff'" (* 1073709055 (0x3fff7fff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x3ffff85e'" (* 1073739870 (0x3ffff85e) *)
- := (Const 1073739870%Z).
-Notation "'0x3ffff85e'" (* 1073739870 (0x3ffff85e) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0x3fffffda'" (* 1073741786 (0x3fffffda) *)
- := (Const 1073741786%Z).
-Notation "'0x3fffffda'" (* 1073741786 (0x3fffffda) *)
- := (Const WO~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~0~1~1~0~1~0).
-Notation "'0x3ffffff6'" (* 1073741814 (0x3ffffff6) *)
- := (Const 1073741814%Z).
-Notation "'0x3ffffff6'" (* 1073741814 (0x3ffffff6) *)
- := (Const WO~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~1~0).
-Notation "'0x3ffffffa'" (* 1073741818 (0x3ffffffa) *)
- := (Const 1073741818%Z).
-Notation "'0x3ffffffa'" (* 1073741818 (0x3ffffffa) *)
- := (Const WO~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~1~0).
-Notation "'0x3ffffffd'" (* 1073741821 (0x3ffffffd) *)
- := (Const 1073741821%Z).
-Notation "'0x3ffffffd'" (* 1073741821 (0x3ffffffd) *)
- := (Const WO~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~1).
-Notation "'0x3ffffffe'" (* 1073741822 (0x3ffffffe) *)
- := (Const 1073741822%Z).
-Notation "'0x3ffffffe'" (* 1073741822 (0x3ffffffe) *)
- := (Const WO~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~0).
-Notation "'0x3fffffff'" (* 1073741823 (0x3fffffff) *)
- := (Const 1073741823%Z).
-Notation "'0x3fffffff'" (* 1073741823 (0x3fffffff) *)
- := (Const WO~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).
-Notation "'0x40000000'" (* 1073741824 (0x40000000) *)
- := (Const 1073741824%Z).
-Notation "'0x40000000'" (* 1073741824 (0x40000000) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000001'" (* 1073741825 (0x40000001) *)
- := (Const 1073741825%Z).
-Notation "'0x40000001'" (* 1073741825 (0x40000001) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x4f72c235'" (* 1332920885 (0x4f72c235) *)
- := (Const 1332920885%Z).
-Notation "'0x4f72c235'" (* 1332920885 (0x4f72c235) *)
- := (Const WO~0~1~0~0~1~1~1~1~0~1~1~1~0~0~1~0~1~1~0~0~0~0~1~0~0~0~1~1~0~1~0~1).
-Notation "'0x684bda13'" (* 1749801491 (0x684bda13) *)
- := (Const 1749801491%Z).
-Notation "'0x684bda13'" (* 1749801491 (0x684bda13) *)
- := (Const WO~0~1~1~0~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~1).
-Notation "'0x7ffefffe'" (* 2147418110 (0x7ffefffe) *)
- := (Const 2147418110%Z).
-Notation "'0x7ffefffe'" (* 2147418110 (0x7ffefffe) *)
- := (Const WO~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x7fffffef'" (* 2147483631 (0x7fffffef) *)
- := (Const 2147483631%Z).
-Notation "'0x7fffffef'" (* 2147483631 (0x7fffffef) *)
- := (Const WO~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~1~1~1).
-Notation "'0x7ffffffa'" (* 2147483642 (0x7ffffffa) *)
- := (Const 2147483642%Z).
-Notation "'0x7ffffffa'" (* 2147483642 (0x7ffffffa) *)
- := (Const WO~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~1~0).
-Notation "'0x7ffffffc'" (* 2147483644 (0x7ffffffc) *)
- := (Const 2147483644%Z).
-Notation "'0x7ffffffc'" (* 2147483644 (0x7ffffffc) *)
- := (Const WO~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~0~0).
-Notation "'0x7ffffffe'" (* 2147483646 (0x7ffffffe) *)
- := (Const 2147483646%Z).
-Notation "'0x7ffffffe'" (* 2147483646 (0x7ffffffe) *)
- := (Const WO~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~0).
-Notation "'0x7fffffff'" (* 2147483647 (0x7fffffff) *)
- := (Const 2147483647%Z).
-Notation "'0x7fffffff'" (* 2147483647 (0x7fffffff) *)
- := (Const WO~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).
-Notation "'0x80000000'" (* 2147483648 (0x80000000) *)
- := (Const 2147483648%Z).
-Notation "'0x80000000'" (* 2147483648 (0x80000000) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000001'" (* 2147483649 (0x80000001) *)
- := (Const 2147483649%Z).
-Notation "'0x80000001'" (* 2147483649 (0x80000001) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xaaaaaaab'" (* 2863311531 (0xaaaaaaab) *)
- := (Const 2863311531%Z).
-Notation "'0xaaaaaaab'" (* 2863311531 (0xaaaaaaab) *)
- := (Const WO~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~1).
-Notation "'0xb0ffffff'" (* 2969567231 (0xb0ffffff) *)
- := (Const 2969567231%Z).
-Notation "'0xb0ffffff'" (* 2969567231 (0xb0ffffff) *)
- := (Const WO~1~0~1~1~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).
-Notation "'0xba2e8ba3'" (* 3123612579 (0xba2e8ba3) *)
- := (Const 3123612579%Z).
-Notation "'0xba2e8ba3'" (* 3123612579 (0xba2e8ba3) *)
- := (Const WO~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~1).
-Notation "'0xc28f5c29'" (* 3264175145 (0xc28f5c29) *)
- := (Const 3264175145%Z).
-Notation "'0xc28f5c29'" (* 3264175145 (0xc28f5c29) *)
- := (Const WO~1~1~0~0~0~0~1~0~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~1).
-Notation "'0xc4ec4ec5'" (* 3303820997 (0xc4ec4ec5) *)
- := (Const 3303820997%Z).
-Notation "'0xc4ec4ec5'" (* 3303820997 (0xc4ec4ec5) *)
- := (Const WO~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~1).
-Notation "'0xcccccccd'" (* 3435973837 (0xcccccccd) *)
- := (Const 3435973837%Z).
-Notation "'0xcccccccd'" (* 3435973837 (0xcccccccd) *)
- := (Const WO~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~1).
-Notation "'0xdcfdcfdd'" (* 3707621341 (0xdcfdcfdd) *)
- := (Const 3707621341%Z).
-Notation "'0xdcfdcfdd'" (* 3707621341 (0xdcfdcfdd) *)
- := (Const WO~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~1).
-Notation "'0xeeeeeeef'" (* 4008636143 (0xeeeeeeef) *)
- := (Const 4008636143%Z).
-Notation "'0xeeeeeeef'" (* 4008636143 (0xeeeeeeef) *)
- := (Const WO~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~1).
-Notation "'0xf0f0f0f1'" (* 4042322161 (0xf0f0f0f1) *)
- := (Const 4042322161%Z).
-Notation "'0xf0f0f0f1'" (* 4042322161 (0xf0f0f0f1) *)
- := (Const WO~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~1).
-Notation "'0xfe14ffff'" (* 4262789119 (0xfe14ffff) *)
- := (Const 4262789119%Z).
-Notation "'0xfe14ffff'" (* 4262789119 (0xfe14ffff) *)
- := (Const WO~1~1~1~1~1~1~1~0~0~0~0~1~0~1~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xffa7ffff'" (* 4289200127 (0xffa7ffff) *)
- := (Const 4289200127%Z).
-Notation "'0xffa7ffff'" (* 4289200127 (0xffa7ffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~0~1~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfffaffff'" (* 4294639615 (0xfffaffff) *)
- := (Const 4294639615%Z).
-Notation "'0xfffaffff'" (* 4294639615 (0xfffaffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfffeffff'" (* 4294901759 (0xfffeffff) *)
- := (Const 4294901759%Z).
-Notation "'0xfffeffff'" (* 4294901759 (0xfffeffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xffffefff'" (* 4294963199 (0xffffefff) *)
- := (Const 4294963199%Z).
-Notation "'0xffffefff'" (* 4294963199 (0xffffefff) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfffffc2f'" (* 4294966319 (0xfffffc2f) *)
- := (Const 4294966319%Z).
-Notation "'0xfffffc2f'" (* 4294966319 (0xfffffc2f) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0xfffffd03'" (* 4294966531 (0xfffffd03) *)
- := (Const 4294966531%Z).
-Notation "'0xfffffd03'" (* 4294966531 (0xfffffd03) *)
- := (Const WO~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~0~0~0~0~0~1~1).
-Notation "'0xfffffdc7'" (* 4294966727 (0xfffffdc7) *)
- := (Const 4294966727%Z).
-Notation "'0xfffffdc7'" (* 4294966727 (0xfffffdc7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0xfffffe1f'" (* 4294966815 (0xfffffe1f) *)
- := (Const 4294966815%Z).
-Notation "'0xfffffe1f'" (* 4294966815 (0xfffffe1f) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~1~1~1~1).
-Notation "'0xfffffe5b'" (* 4294966875 (0xfffffe5b) *)
- := (Const 4294966875%Z).
-Notation "'0xfffffe5b'" (* 4294966875 (0xfffffe5b) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1).
-Notation "'0xfffffec3'" (* 4294966979 (0xfffffec3) *)
- := (Const 4294966979%Z).
-Notation "'0xfffffec3'" (* 4294966979 (0xfffffec3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0xffffff43'" (* 4294967107 (0xffffff43) *)
- := (Const 4294967107%Z).
-Notation "'0xffffff43'" (* 4294967107 (0xffffff43) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0xffffff45'" (* 4294967109 (0xffffff45) *)
- := (Const 4294967109%Z).
-Notation "'0xffffff45'" (* 4294967109 (0xffffff45) *)
- := (Const WO~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~0~0~1~0~1).
-Notation "'0xffffff8b'" (* 4294967179 (0xffffff8b) *)
- := (Const 4294967179%Z).
-Notation "'0xffffff8b'" (* 4294967179 (0xffffff8b) *)
- := (Const WO~1~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~0~0~1~0~1~1).
-Notation "'0xffffff97'" (* 4294967191 (0xffffff97) *)
- := (Const 4294967191%Z).
-Notation "'0xffffff97'" (* 4294967191 (0xffffff97) *)
- := (Const WO~1~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~0~1~0~1~1~1).
-Notation "'0xffffffde'" (* 4294967262 (0xffffffde) *)
- := (Const 4294967262%Z).
-Notation "'0xffffffde'" (* 4294967262 (0xffffffde) *)
- := (Const WO~1~1~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~1~1~0).
-Notation "'0xffffffdf'" (* 4294967263 (0xffffffdf) *)
- := (Const 4294967263%Z).
-Notation "'0xffffffdf'" (* 4294967263 (0xffffffdf) *)
- := (Const WO~1~1~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~1~1~1).
-Notation "'0xffffffe1'" (* 4294967265 (0xffffffe1) *)
- := (Const 4294967265%Z).
-Notation "'0xffffffe1'" (* 4294967265 (0xffffffe1) *)
- := (Const WO~1~1~1~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~0~0~0~1).
-Notation "'0xffffffe3'" (* 4294967267 (0xffffffe3) *)
- := (Const 4294967267%Z).
-Notation "'0xffffffe3'" (* 4294967267 (0xffffffe3) *)
- := (Const WO~1~1~1~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~0~0~1~1).
-Notation "'0xffffffe5'" (* 4294967269 (0xffffffe5) *)
- := (Const 4294967269%Z).
-Notation "'0xffffffe5'" (* 4294967269 (0xffffffe5) *)
- := (Const WO~1~1~1~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~0~1~0~1).
-Notation "'0xffffffe7'" (* 4294967271 (0xffffffe7) *)
- := (Const 4294967271%Z).
-Notation "'0xffffffe7'" (* 4294967271 (0xffffffe7) *)
- := (Const WO~1~1~1~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~0~1~1~1).
-Notation "'0xffffffeb'" (* 4294967275 (0xffffffeb) *)
- := (Const 4294967275%Z).
-Notation "'0xffffffeb'" (* 4294967275 (0xffffffeb) *)
- := (Const WO~1~1~1~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~1~1).
-Notation "'0xffffffed'" (* 4294967277 (0xffffffed) *)
- := (Const 4294967277%Z).
-Notation "'0xffffffed'" (* 4294967277 (0xffffffed) *)
- := (Const WO~1~1~1~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).
-Notation "'0xffffffef'" (* 4294967279 (0xffffffef) *)
- := (Const 4294967279%Z).
-Notation "'0xffffffef'" (* 4294967279 (0xffffffef) *)
- := (Const WO~1~1~1~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~1~1).
-Notation "'0xfffffff1'" (* 4294967281 (0xfffffff1) *)
- := (Const 4294967281%Z).
-Notation "'0xfffffff1'" (* 4294967281 (0xfffffff1) *)
- := (Const WO~1~1~1~1~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~0~0~1).
-Notation "'0xfffffff3'" (* 4294967283 (0xfffffff3) *)
- := (Const 4294967283%Z).
-Notation "'0xfffffff3'" (* 4294967283 (0xfffffff3) *)
- := (Const WO~1~1~1~1~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~0~1~1).
-Notation "'0xfffffff5'" (* 4294967285 (0xfffffff5) *)
- := (Const 4294967285%Z).
-Notation "'0xfffffff5'" (* 4294967285 (0xfffffff5) *)
- := (Const WO~1~1~1~1~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~1).
-Notation "'0xfffffff7'" (* 4294967287 (0xfffffff7) *)
- := (Const 4294967287%Z).
-Notation "'0xfffffff7'" (* 4294967287 (0xfffffff7) *)
- := (Const WO~1~1~1~1~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~1).
-Notation "'0xfffffffb'" (* 4294967291 (0xfffffffb) *)
- := (Const 4294967291%Z).
-Notation "'0xfffffffb'" (* 4294967291 (0xfffffffb) *)
- := (Const WO~1~1~1~1~1~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).
-Notation "'0xfffffffd'" (* 4294967293 (0xfffffffd) *)
- := (Const 4294967293%Z).
-Notation "'0xfffffffd'" (* 4294967293 (0xfffffffd) *)
- := (Const WO~1~1~1~1~1~1~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).
-Notation "'0xfffffffe'" (* 4294967294 (0xfffffffe) *)
- := (Const 4294967294%Z).
-Notation "'0xfffffffe'" (* 4294967294 (0xfffffffe) *)
- := (Const WO~1~1~1~1~1~1~1~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 "'0xffffffff'" (* 4294967295 (0xffffffff) *)
- := (Const 4294967295%Z).
-Notation "'0xffffffff'" (* 4294967295 (0xffffffff) *)
- := (Const WO~1~1~1~1~1~1~1~1~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 "'0x100000000'" (* 4294967296 (0x100000000) *)
- := (Const 4294967296%Z).
-Notation "'0x100000000'" (* 4294967296 (0x100000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000001'" (* 4294967297 (0x100000001) *)
- := (Const 4294967297%Z).
-Notation "'0x100000001'" (* 4294967297 (0x100000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1000003d1'" (* 4294968273 (0x1000003d1) *)
- := (Const 4294968273%Z).
-Notation "'0x1000003d1'" (* 4294968273 (0x1000003d1) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~0~1~0~0~0~1).
-Notation "'0x1ffffffdf'" (* 8589934559 (0x1ffffffdf) *)
- := (Const 8589934559%Z).
-Notation "'0x1ffffffdf'" (* 8589934559 (0x1ffffffdf) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1).
-Notation "'0x1ffffffe7'" (* 8589934567 (0x1ffffffe7) *)
- := (Const 8589934567%Z).
-Notation "'0x1ffffffe7'" (* 8589934567 (0x1ffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0x1ffffffef'" (* 8589934575 (0x1ffffffef) *)
- := (Const 8589934575%Z).
-Notation "'0x1ffffffef'" (* 8589934575 (0x1ffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0x1fffffffb'" (* 8589934587 (0x1fffffffb) *)
- := (Const 8589934587%Z).
-Notation "'0x1fffffffb'" (* 8589934587 (0x1fffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0x1fffffffe'" (* 8589934590 (0x1fffffffe) *)
- := (Const 8589934590%Z).
-Notation "'0x1fffffffe'" (* 8589934590 (0x1fffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x1ffffffff'" (* 8589934591 (0x1ffffffff) *)
- := (Const 8589934591%Z).
-Notation "'0x1ffffffff'" (* 8589934591 (0x1ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x200000000'" (* 8589934592 (0x200000000) *)
- := (Const 8589934592%Z).
-Notation "'0x200000000'" (* 8589934592 (0x200000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000001'" (* 8589934593 (0x200000001) *)
- := (Const 8589934593%Z).
-Notation "'0x200000001'" (* 8589934593 (0x200000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffce'" (* 17179869134 (0x3ffffffce) *)
- := (Const 17179869134%Z).
-Notation "'0x3ffffffce'" (* 17179869134 (0x3ffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0x3fffffff6'" (* 17179869174 (0x3fffffff6) *)
- := (Const 17179869174%Z).
-Notation "'0x3fffffff6'" (* 17179869174 (0x3fffffff6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0x3fffffffe'" (* 17179869182 (0x3fffffffe) *)
- := (Const 17179869182%Z).
-Notation "'0x3fffffffe'" (* 17179869182 (0x3fffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x3ffffffff'" (* 17179869183 (0x3ffffffff) *)
- := (Const 17179869183%Z).
-Notation "'0x3ffffffff'" (* 17179869183 (0x3ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x400000000'" (* 17179869184 (0x400000000) *)
- := (Const 17179869184%Z).
-Notation "'0x400000000'" (* 17179869184 (0x400000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000001'" (* 17179869185 (0x400000001) *)
- := (Const 17179869185%Z).
-Notation "'0x400000001'" (* 17179869185 (0x400000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffe5'" (* 34359738341 (0x7ffffffe5) *)
- := (Const 34359738341%Z).
-Notation "'0x7ffffffe5'" (* 34359738341 (0x7ffffffe5) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1).
-Notation "'0x7ffffffed'" (* 34359738349 (0x7ffffffed) *)
- := (Const 34359738349%Z).
-Notation "'0x7ffffffed'" (* 34359738349 (0x7ffffffed) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0~1).
-Notation "'0x7fffffff3'" (* 34359738355 (0x7fffffff3) *)
- := (Const 34359738355%Z).
-Notation "'0x7fffffff3'" (* 34359738355 (0x7fffffff3) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1).
-Notation "'0x7fffffffe'" (* 34359738366 (0x7fffffffe) *)
- := (Const 34359738366%Z).
-Notation "'0x7fffffffe'" (* 34359738366 (0x7fffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x7ffffffff'" (* 34359738367 (0x7ffffffff) *)
- := (Const 34359738367%Z).
-Notation "'0x7ffffffff'" (* 34359738367 (0x7ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x800000000'" (* 34359738368 (0x800000000) *)
- := (Const 34359738368%Z).
-Notation "'0x800000000'" (* 34359738368 (0x800000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000001'" (* 34359738369 (0x800000001) *)
- := (Const 34359738369%Z).
-Notation "'0x800000001'" (* 34359738369 (0x800000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffca'" (* 68719476682 (0xfffffffca) *)
- := (Const 68719476682%Z).
-Notation "'0xfffffffca'" (* 68719476682 (0xfffffffca) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1~0).
-Notation "'0xfffffffe3'" (* 68719476707 (0xfffffffe3) *)
- := (Const 68719476707%Z).
-Notation "'0xfffffffe3'" (* 68719476707 (0xfffffffe3) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1~1).
-Notation "'0xfffffffe6'" (* 68719476710 (0xfffffffe6) *)
- := (Const 68719476710%Z).
-Notation "'0xfffffffe6'" (* 68719476710 (0xfffffffe6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~0).
-Notation "'0xffffffff7'" (* 68719476727 (0xffffffff7) *)
- := (Const 68719476727%Z).
-Notation "'0xffffffff7'" (* 68719476727 (0xffffffff7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1).
-Notation "'0xffffffffd'" (* 68719476733 (0xffffffffd) *)
- := (Const 68719476733%Z).
-Notation "'0xffffffffd'" (* 68719476733 (0xffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1).
-Notation "'0xffffffffe'" (* 68719476734 (0xffffffffe) *)
- := (Const 68719476734%Z).
-Notation "'0xffffffffe'" (* 68719476734 (0xffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0xfffffffff'" (* 68719476735 (0xfffffffff) *)
- := (Const 68719476735%Z).
-Notation "'0xfffffffff'" (* 68719476735 (0xfffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x1000000000'" (* 68719476736 (0x1000000000) *)
- := (Const 68719476736%Z).
-Notation "'0x1000000000'" (* 68719476736 (0x1000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000001'" (* 68719476737 (0x1000000001) *)
- := (Const 68719476737%Z).
-Notation "'0x1000000001'" (* 68719476737 (0x1000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1efffffc2f'" (* 133143985199 (0x1efffffc2f) *)
- := (Const 133143985199%Z).
-Notation "'0x1efffffc2f'" (* 133143985199 (0x1efffffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0x1fffe00000'" (* 137436856320 (0x1fffe00000) *)
- := (Const 137436856320%Z).
-Notation "'0x1fffe00000'" (* 137436856320 (0x1fffe00000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1ffffffe1f'" (* 137438952991 (0x1ffffffe1f) *)
- := (Const 137438952991%Z).
-Notation "'0x1ffffffe1f'" (* 137438952991 (0x1ffffffe1f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~0~1~1~1~1~1).
-Notation "'0x1fffffff45'" (* 137438953285 (0x1fffffff45) *)
- := (Const 137438953285%Z).
-Notation "'0x1fffffff45'" (* 137438953285 (0x1fffffff45) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0~0~0~1~0~1).
-Notation "'0x1fffffff8b'" (* 137438953355 (0x1fffffff8b) *)
- := (Const 137438953355%Z).
-Notation "'0x1fffffff8b'" (* 137438953355 (0x1fffffff8b) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1~0~1~1).
-Notation "'0x1fffffffee'" (* 137438953454 (0x1fffffffee) *)
- := (Const 137438953454%Z).
-Notation "'0x1fffffffee'" (* 137438953454 (0x1fffffffee) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~0).
-Notation "'0x1ffffffffe'" (* 137438953470 (0x1ffffffffe) *)
- := (Const 137438953470%Z).
-Notation "'0x1ffffffffe'" (* 137438953470 (0x1ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x1fffffffff'" (* 137438953471 (0x1fffffffff) *)
- := (Const 137438953471%Z).
-Notation "'0x1fffffffff'" (* 137438953471 (0x1fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x2000000000'" (* 137438953472 (0x2000000000) *)
- := (Const 137438953472%Z).
-Notation "'0x2000000000'" (* 137438953472 (0x2000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000001'" (* 137438953473 (0x2000000001) *)
- := (Const 137438953473%Z).
-Notation "'0x2000000001'" (* 137438953473 (0x2000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3dfffff85e'" (* 266287970398 (0x3dfffff85e) *)
- := (Const 266287970398%Z).
-Notation "'0x3dfffff85e'" (* 266287970398 (0x3dfffff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0x3ffdffffff'" (* 274844352511 (0x3ffdffffff) *)
- := (Const 274844352511%Z).
-Notation "'0x3ffdffffff'" (* 274844352511 (0x3ffdffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~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 "'0x3fffffefff'" (* 274877902847 (0x3fffffefff) *)
- := (Const 274877902847%Z).
-Notation "'0x3fffffefff'" (* 274877902847 (0x3fffffefff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x3fffffffe7'" (* 274877906919 (0x3fffffffe7) *)
- := (Const 274877906919%Z).
-Notation "'0x3fffffffe7'" (* 274877906919 (0x3fffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0x3fffffffef'" (* 274877906927 (0x3fffffffef) *)
- := (Const 274877906927%Z).
-Notation "'0x3fffffffef'" (* 274877906927 (0x3fffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0x3ffffffff5'" (* 274877906933 (0x3ffffffff5) *)
- := (Const 274877906933%Z).
-Notation "'0x3ffffffff5'" (* 274877906933 (0x3ffffffff5) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0~1).
-Notation "'0x3ffffffffb'" (* 274877906939 (0x3ffffffffb) *)
- := (Const 274877906939%Z).
-Notation "'0x3ffffffffb'" (* 274877906939 (0x3ffffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0x3ffffffffd'" (* 274877906941 (0x3ffffffffd) *)
- := (Const 274877906941%Z).
-Notation "'0x3ffffffffd'" (* 274877906941 (0x3ffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1).
-Notation "'0x3ffffffffe'" (* 274877906942 (0x3ffffffffe) *)
- := (Const 274877906942%Z).
-Notation "'0x3ffffffffe'" (* 274877906942 (0x3ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x3fffffffff'" (* 274877906943 (0x3fffffffff) *)
- := (Const 274877906943%Z).
-Notation "'0x3fffffffff'" (* 274877906943 (0x3fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x4000000000'" (* 274877906944 (0x4000000000) *)
- := (Const 274877906944%Z).
-Notation "'0x4000000000'" (* 274877906944 (0x4000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000001'" (* 274877906945 (0x4000000001) *)
- := (Const 274877906945%Z).
-Notation "'0x4000000001'" (* 274877906945 (0x4000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffbfffffe'" (* 549688705022 (0x7ffbfffffe) *)
- := (Const 549688705022%Z).
-Notation "'0x7ffbfffffe'" (* 549688705022 (0x7ffbfffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~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 "'0x7ffff80000'" (* 549755289600 (0x7ffff80000) *)
- := (Const 549755289600%Z).
-Notation "'0x7ffff80000'" (* 549755289600 (0x7ffff80000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x7fffffff97'" (* 549755813783 (0x7fffffff97) *)
- := (Const 549755813783%Z).
-Notation "'0x7fffffff97'" (* 549755813783 (0x7fffffff97) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1~1~1).
-Notation "'0x7fffffffce'" (* 549755813838 (0x7fffffffce) *)
- := (Const 549755813838%Z).
-Notation "'0x7fffffffce'" (* 549755813838 (0x7fffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0x7fffffffde'" (* 549755813854 (0x7fffffffde) *)
- := (Const 549755813854%Z).
-Notation "'0x7fffffffde'" (* 549755813854 (0x7fffffffde) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0x7fffffffdf'" (* 549755813855 (0x7fffffffdf) *)
- := (Const 549755813855%Z).
-Notation "'0x7fffffffdf'" (* 549755813855 (0x7fffffffdf) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1).
-Notation "'0x7fffffffea'" (* 549755813866 (0x7fffffffea) *)
- := (Const 549755813866%Z).
-Notation "'0x7fffffffea'" (* 549755813866 (0x7fffffffea) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0~1~0).
-Notation "'0x7fffffffed'" (* 549755813869 (0x7fffffffed) *)
- := (Const 549755813869%Z).
-Notation "'0x7fffffffed'" (* 549755813869 (0x7fffffffed) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0~1).
-Notation "'0x7ffffffff6'" (* 549755813878 (0x7ffffffff6) *)
- := (Const 549755813878%Z).
-Notation "'0x7ffffffff6'" (* 549755813878 (0x7ffffffff6) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0x7ffffffffa'" (* 549755813882 (0x7ffffffffa) *)
- := (Const 549755813882%Z).
-Notation "'0x7ffffffffa'" (* 549755813882 (0x7ffffffffa) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0~1~0).
-Notation "'0x7ffffffffe'" (* 549755813886 (0x7ffffffffe) *)
- := (Const 549755813886%Z).
-Notation "'0x7ffffffffe'" (* 549755813886 (0x7ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x7fffffffff'" (* 549755813887 (0x7fffffffff) *)
- := (Const 549755813887%Z).
-Notation "'0x7fffffffff'" (* 549755813887 (0x7fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x8000000000'" (* 549755813888 (0x8000000000) *)
- := (Const 549755813888%Z).
-Notation "'0x8000000000'" (* 549755813888 (0x8000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000001'" (* 549755813889 (0x8000000001) *)
- := (Const 549755813889%Z).
-Notation "'0x8000000001'" (* 549755813889 (0x8000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffff2e'" (* 1099511627566 (0xffffffff2e) *)
- := (Const 1099511627566%Z).
-Notation "'0xffffffff2e'" (* 1099511627566 (0xffffffff2e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~0~1~0~1~1~1~0).
-Notation "'0xffffffffbe'" (* 1099511627710 (0xffffffffbe) *)
- := (Const 1099511627710%Z).
-Notation "'0xffffffffbe'" (* 1099511627710 (0xffffffffbe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1~0).
-Notation "'0xffffffffda'" (* 1099511627738 (0xffffffffda) *)
- := (Const 1099511627738%Z).
-Notation "'0xffffffffda'" (* 1099511627738 (0xffffffffda) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~0~1~0).
-Notation "'0xffffffffef'" (* 1099511627759 (0xffffffffef) *)
- := (Const 1099511627759%Z).
-Notation "'0xffffffffef'" (* 1099511627759 (0xffffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0xfffffffff1'" (* 1099511627761 (0xfffffffff1) *)
- := (Const 1099511627761%Z).
-Notation "'0xfffffffff1'" (* 1099511627761 (0xfffffffff1) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1).
-Notation "'0xfffffffffe'" (* 1099511627774 (0xfffffffffe) *)
- := (Const 1099511627774%Z).
-Notation "'0xfffffffffe'" (* 1099511627774 (0xfffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0xffffffffff'" (* 1099511627775 (0xffffffffff) *)
- := (Const 1099511627775%Z).
-Notation "'0xffffffffff'" (* 1099511627775 (0xffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x10000000000'" (* 1099511627776 (0x10000000000) *)
- := (Const 1099511627776%Z).
-Notation "'0x10000000000'" (* 1099511627776 (0x10000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000001'" (* 1099511627777 (0x10000000001) *)
- := (Const 1099511627777%Z).
-Notation "'0x10000000001'" (* 1099511627777 (0x10000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x14bffffffff'" (* 1425929142271 (0x14bffffffff) *)
- := (Const 1425929142271%Z).
-Notation "'0x14bffffffff'" (* 1425929142271 (0x14bffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~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).
-Notation "'0x1ffffffffde'" (* 2199023255518 (0x1ffffffffde) *)
- := (Const 2199023255518%Z).
-Notation "'0x1ffffffffde'" (* 2199023255518 (0x1ffffffffde) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0x1ffffffffe2'" (* 2199023255522 (0x1ffffffffe2) *)
- := (Const 2199023255522%Z).
-Notation "'0x1ffffffffe2'" (* 2199023255522 (0x1ffffffffe2) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0~0~0~1~0).
-Notation "'0x1fffffffff7'" (* 2199023255543 (0x1fffffffff7) *)
- := (Const 2199023255543%Z).
-Notation "'0x1fffffffff7'" (* 2199023255543 (0x1fffffffff7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0~1~1~1).
-Notation "'0x1fffffffffe'" (* 2199023255550 (0x1fffffffffe) *)
- := (Const 2199023255550%Z).
-Notation "'0x1fffffffffe'" (* 2199023255550 (0x1fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x1ffffffffff'" (* 2199023255551 (0x1ffffffffff) *)
- := (Const 2199023255551%Z).
-Notation "'0x1ffffffffff'" (* 2199023255551 (0x1ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x20000000000'" (* 2199023255552 (0x20000000000) *)
- := (Const 2199023255552%Z).
-Notation "'0x20000000000'" (* 2199023255552 (0x20000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000001'" (* 2199023255553 (0x20000000001) *)
- := (Const 2199023255553%Z).
-Notation "'0x20000000001'" (* 2199023255553 (0x20000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3f80fffffff'" (* 4363955208191 (0x3f80fffffff) *)
- := (Const 4363955208191%Z).
-Notation "'0x3f80fffffff'" (* 4363955208191 (0x3f80fffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~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).
-Notation "'0x3fffffffc00'" (* 4398046510080 (0x3fffffffc00) *)
- := (Const 4398046510080%Z).
-Notation "'0x3fffffffc00'" (* 4398046510080 (0x3fffffffc00) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x3ffffffffe7'" (* 4398046511079 (0x3ffffffffe7) *)
- := (Const 4398046511079%Z).
-Notation "'0x3ffffffffe7'" (* 4398046511079 (0x3ffffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0x3ffffffffee'" (* 4398046511086 (0x3ffffffffee) *)
- := (Const 4398046511086%Z).
-Notation "'0x3ffffffffee'" (* 4398046511086 (0x3ffffffffee) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~1~1~1~0).
-Notation "'0x3fffffffff7'" (* 4398046511095 (0x3fffffffff7) *)
- := (Const 4398046511095%Z).
-Notation "'0x3fffffffff7'" (* 4398046511095 (0x3fffffffff7) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~1~1~1).
-Notation "'0x3fffffffffb'" (* 4398046511099 (0x3fffffffffb) *)
- := (Const 4398046511099%Z).
-Notation "'0x3fffffffffb'" (* 4398046511099 (0x3fffffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0x3fffffffffe'" (* 4398046511102 (0x3fffffffffe) *)
- := (Const 4398046511102%Z).
-Notation "'0x3fffffffffe'" (* 4398046511102 (0x3fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x3ffffffffff'" (* 4398046511103 (0x3ffffffffff) *)
- := (Const 4398046511103%Z).
-Notation "'0x3ffffffffff'" (* 4398046511103 (0x3ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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).
-Notation "'0x40000000000'" (* 4398046511104 (0x40000000000) *)
- := (Const 4398046511104%Z).
-Notation "'0x40000000000'" (* 4398046511104 (0x40000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000001'" (* 4398046511105 (0x40000000001) *)
- := (Const 4398046511105%Z).
-Notation "'0x40000000001'" (* 4398046511105 (0x40000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7f01ffffffe'" (* 8727910416382 (0x7f01ffffffe) *)
- := (Const 8727910416382%Z).
-Notation "'0x7f01ffffffe'" (* 8727910416382 (0x7f01ffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~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~0).
-Notation "'0x7fefffffc2f'" (* 8791798053935 (0x7fefffffc2f) *)
- := (Const 8791798053935%Z).
-Notation "'0x7fefffffc2f'" (* 8791798053935 (0x7fefffffc2f) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0x7ffffdfffff'" (* 8796090925055 (0x7ffffdfffff) *)
- := (Const 8796090925055%Z).
-Notation "'0x7ffffdfffff'" (* 8796090925055 (0x7ffffdfffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x7fffffffd03'" (* 8796093021443 (0x7fffffffd03) *)
- := (Const 8796093021443%Z).
-Notation "'0x7fffffffd03'" (* 8796093021443 (0x7fffffffd03) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~0~0~0~0~0~0~1~1).
-Notation "'0x7ffffffff43'" (* 8796093022019 (0x7ffffffff43) *)
- := (Const 8796093022019%Z).
-Notation "'0x7ffffffff43'" (* 8796093022019 (0x7ffffffff43) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~0~0~0~0~1~1).
-Notation "'0x7ffffffffce'" (* 8796093022158 (0x7ffffffffce) *)
- := (Const 8796093022158%Z).
-Notation "'0x7ffffffffce'" (* 8796093022158 (0x7ffffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0x7ffffffffe3'" (* 8796093022179 (0x7ffffffffe3) *)
- := (Const 8796093022179%Z).
-Notation "'0x7ffffffffe3'" (* 8796093022179 (0x7ffffffffe3) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~0~1~1).
-Notation "'0x7ffffffffe7'" (* 8796093022183 (0x7ffffffffe7) *)
- := (Const 8796093022183%Z).
-Notation "'0x7ffffffffe7'" (* 8796093022183 (0x7ffffffffe7) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~1~1~1).
-Notation "'0x7ffffffffed'" (* 8796093022189 (0x7ffffffffed) *)
- := (Const 8796093022189%Z).
-Notation "'0x7ffffffffed'" (* 8796093022189 (0x7ffffffffed) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~0~1).
-Notation "'0x7ffffffffee'" (* 8796093022190 (0x7ffffffffee) *)
- := (Const 8796093022190%Z).
-Notation "'0x7ffffffffee'" (* 8796093022190 (0x7ffffffffee) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~1~0).
-Notation "'0x7fffffffff1'" (* 8796093022193 (0x7fffffffff1) *)
- := (Const 8796093022193%Z).
-Notation "'0x7fffffffff1'" (* 8796093022193 (0x7fffffffff1) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~0~0~1).
-Notation "'0x7fffffffff6'" (* 8796093022198 (0x7fffffffff6) *)
- := (Const 8796093022198%Z).
-Notation "'0x7fffffffff6'" (* 8796093022198 (0x7fffffffff6) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0x7fffffffffd'" (* 8796093022205 (0x7fffffffffd) *)
- := (Const 8796093022205%Z).
-Notation "'0x7fffffffffd'" (* 8796093022205 (0x7fffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0~1).
-Notation "'0x7fffffffffe'" (* 8796093022206 (0x7fffffffffe) *)
- := (Const 8796093022206%Z).
-Notation "'0x7fffffffffe'" (* 8796093022206 (0x7fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~0).
-Notation "'0x7ffffffffff'" (* 8796093022207 (0x7ffffffffff) *)
- := (Const 8796093022207%Z).
-Notation "'0x7ffffffffff'" (* 8796093022207 (0x7ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~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).
-Notation "'0x80000000000'" (* 8796093022208 (0x80000000000) *)
- := (Const 8796093022208%Z).
-Notation "'0x80000000000'" (* 8796093022208 (0x80000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000001'" (* 8796093022209 (0x80000000001) *)
- := (Const 8796093022209%Z).
-Notation "'0x80000000001'" (* 8796093022209 (0x80000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffdfffff85e'" (* 17583596107870 (0xffdfffff85e) *)
- := (Const 17583596107870%Z).
-Notation "'0xffdfffff85e'" (* 17583596107870 (0xffdfffff85e) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0xfffffbffffe'" (* 17592181850110 (0xfffffbffffe) *)
- := (Const 17592181850110%Z).
-Notation "'0xfffffbffffe'" (* 17592181850110 (0xfffffbffffe) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0xffffffffa06'" (* 17592186042886 (0xffffffffa06) *)
- := (Const 17592186042886%Z).
-Notation "'0xffffffffa06'" (* 17592186042886 (0xffffffffa06) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0xffffffffe86'" (* 17592186044038 (0xffffffffe86) *)
- := (Const 17592186044038%Z).
-Notation "'0xffffffffe86'" (* 17592186044038 (0xffffffffe86) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~0~0~0~0~1~1~0).
-Notation "'0xfffffffffc6'" (* 17592186044358 (0xfffffffffc6) *)
- := (Const 17592186044358%Z).
-Notation "'0xfffffffffc6'" (* 17592186044358 (0xfffffffffc6) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~0~0~1~1~0).
-Notation "'0xfffffffffce'" (* 17592186044366 (0xfffffffffce) *)
- := (Const 17592186044366%Z).
-Notation "'0xfffffffffce'" (* 17592186044366 (0xfffffffffce) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~0~1~1~1~0).
-Notation "'0xfffffffffda'" (* 17592186044378 (0xfffffffffda) *)
- := (Const 17592186044378%Z).
-Notation "'0xfffffffffda'" (* 17592186044378 (0xfffffffffda) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1~0~1~0).
-Notation "'0xfffffffffef'" (* 17592186044399 (0xfffffffffef) *)
- := (Const 17592186044399%Z).
-Notation "'0xfffffffffef'" (* 17592186044399 (0xfffffffffef) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0xffffffffffa'" (* 17592186044410 (0xffffffffffa) *)
- := (Const 17592186044410%Z).
-Notation "'0xffffffffffa'" (* 17592186044410 (0xffffffffffa) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~0).
-Notation "'0xffffffffffb'" (* 17592186044411 (0xffffffffffb) *)
- := (Const 17592186044411%Z).
-Notation "'0xffffffffffb'" (* 17592186044411 (0xffffffffffb) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1~1).
-Notation "'0xffffffffffd'" (* 17592186044413 (0xffffffffffd) *)
- := (Const 17592186044413%Z).
-Notation "'0xffffffffffd'" (* 17592186044413 (0xffffffffffd) *)
- := (Const WO~0~0~0~0~0~0~0~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~0~1).
-Notation "'0xffffffffffe'" (* 17592186044414 (0xffffffffffe) *)
- := (Const 17592186044414%Z).
-Notation "'0xffffffffffe'" (* 17592186044414 (0xffffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~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~0).
-Notation "'0xfffffffffff'" (* 17592186044415 (0xfffffffffff) *)
- := (Const 17592186044415%Z).
-Notation "'0xfffffffffff'" (* 17592186044415 (0xfffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~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).
-Notation "'0x100000000000'" (* 17592186044416 (0x100000000000) *)
- := (Const 17592186044416%Z).
-Notation "'0x100000000000'" (* 17592186044416 (0x100000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000001'" (* 17592186044417 (0x100000000001) *)
- := (Const 17592186044417%Z).
-Notation "'0x100000000001'" (* 17592186044417 (0x100000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffff8b'" (* 35184372088715 (0x1fffffffff8b) *)
- := (Const 35184372088715%Z).
-Notation "'0x1fffffffff8b'" (* 35184372088715 (0x1fffffffff8b) *)
- := (Const WO~0~0~0~0~0~0~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~0~0~0~1~0~1~1).
-Notation "'0x1fffffffffc0'" (* 35184372088768 (0x1fffffffffc0) *)
- := (Const 35184372088768%Z).
-Notation "'0x1fffffffffc0'" (* 35184372088768 (0x1fffffffffc0) *)
- := (Const WO~0~0~0~0~0~0~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~0~0~0~0~0~0).
-Notation "'0x1fffffffffde'" (* 35184372088798 (0x1fffffffffde) *)
- := (Const 35184372088798%Z).
-Notation "'0x1fffffffffde'" (* 35184372088798 (0x1fffffffffde) *)
- := (Const WO~0~0~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0x1ffffffffff6'" (* 35184372088822 (0x1ffffffffff6) *)
- := (Const 35184372088822%Z).
-Notation "'0x1ffffffffff6'" (* 35184372088822 (0x1ffffffffff6) *)
- := (Const WO~0~0~0~0~0~0~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~0~1~1~0).
-Notation "'0x1ffffffffffa'" (* 35184372088826 (0x1ffffffffffa) *)
- := (Const 35184372088826%Z).
-Notation "'0x1ffffffffffa'" (* 35184372088826 (0x1ffffffffffa) *)
- := (Const WO~0~0~0~0~0~0~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~0~1~0).
-Notation "'0x1ffffffffffd'" (* 35184372088829 (0x1ffffffffffd) *)
- := (Const 35184372088829%Z).
-Notation "'0x1ffffffffffd'" (* 35184372088829 (0x1ffffffffffd) *)
- := (Const WO~0~0~0~0~0~0~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~0~1).
-Notation "'0x1ffffffffffe'" (* 35184372088830 (0x1ffffffffffe) *)
- := (Const 35184372088830%Z).
-Notation "'0x1ffffffffffe'" (* 35184372088830 (0x1ffffffffffe) *)
- := (Const WO~0~0~0~0~0~0~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~0).
-Notation "'0x1fffffffffff'" (* 35184372088831 (0x1fffffffffff) *)
- := (Const 35184372088831%Z).
-Notation "'0x1fffffffffff'" (* 35184372088831 (0x1fffffffffff) *)
- := (Const WO~0~0~0~0~0~0~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).
-Notation "'0x200000000000'" (* 35184372088832 (0x200000000000) *)
- := (Const 35184372088832%Z).
-Notation "'0x200000000000'" (* 35184372088832 (0x200000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000001'" (* 35184372088833 (0x200000000001) *)
- := (Const 35184372088833%Z).
-Notation "'0x200000000001'" (* 35184372088833 (0x200000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffff7fffff'" (* 70368735789055 (0x3fffff7fffff) *)
- := (Const 70368735789055%Z).
-Notation "'0x3fffff7fffff'" (* 70368735789055 (0x3fffff7fffff) *)
- := (Const WO~0~0~0~0~0~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~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 "'0x3fffffffff16'" (* 70368744177430 (0x3fffffffff16) *)
- := (Const 70368744177430%Z).
-Notation "'0x3fffffffff16'" (* 70368744177430 (0x3fffffffff16) *)
- := (Const WO~0~0~0~0~0~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~0~0~0~1~0~1~1~0).
-Notation "'0x3fffffffffe5'" (* 70368744177637 (0x3fffffffffe5) *)
- := (Const 70368744177637%Z).
-Notation "'0x3fffffffffe5'" (* 70368744177637 (0x3fffffffffe5) *)
- := (Const WO~0~0~0~0~0~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~0~0~1~0~1).
-Notation "'0x3fffffffffef'" (* 70368744177647 (0x3fffffffffef) *)
- := (Const 70368744177647%Z).
-Notation "'0x3fffffffffef'" (* 70368744177647 (0x3fffffffffef) *)
- := (Const WO~0~0~0~0~0~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~0~1~1~1~1).
-Notation "'0x3ffffffffff3'" (* 70368744177651 (0x3ffffffffff3) *)
- := (Const 70368744177651%Z).
-Notation "'0x3ffffffffff3'" (* 70368744177651 (0x3ffffffffff3) *)
- := (Const WO~0~0~0~0~0~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~0~0~1~1).
-Notation "'0x3ffffffffff7'" (* 70368744177655 (0x3ffffffffff7) *)
- := (Const 70368744177655%Z).
-Notation "'0x3ffffffffff7'" (* 70368744177655 (0x3ffffffffff7) *)
- := (Const WO~0~0~0~0~0~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~0~1~1~1).
-Notation "'0x3ffffffffffa'" (* 70368744177658 (0x3ffffffffffa) *)
- := (Const 70368744177658%Z).
-Notation "'0x3ffffffffffa'" (* 70368744177658 (0x3ffffffffffa) *)
- := (Const WO~0~0~0~0~0~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~0~1~0).
-Notation "'0x3ffffffffffb'" (* 70368744177659 (0x3ffffffffffb) *)
- := (Const 70368744177659%Z).
-Notation "'0x3ffffffffffb'" (* 70368744177659 (0x3ffffffffffb) *)
- := (Const WO~0~0~0~0~0~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~0~1~1).
-Notation "'0x3ffffffffffc'" (* 70368744177660 (0x3ffffffffffc) *)
- := (Const 70368744177660%Z).
-Notation "'0x3ffffffffffc'" (* 70368744177660 (0x3ffffffffffc) *)
- := (Const WO~0~0~0~0~0~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~0~0).
-Notation "'0x3ffffffffffd'" (* 70368744177661 (0x3ffffffffffd) *)
- := (Const 70368744177661%Z).
-Notation "'0x3ffffffffffd'" (* 70368744177661 (0x3ffffffffffd) *)
- := (Const WO~0~0~0~0~0~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~0~1).
-Notation "'0x3ffffffffffe'" (* 70368744177662 (0x3ffffffffffe) *)
- := (Const 70368744177662%Z).
-Notation "'0x3ffffffffffe'" (* 70368744177662 (0x3ffffffffffe) *)
- := (Const WO~0~0~0~0~0~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~0).
-Notation "'0x3fffffffffff'" (* 70368744177663 (0x3fffffffffff) *)
- := (Const 70368744177663%Z).
-Notation "'0x3fffffffffff'" (* 70368744177663 (0x3fffffffffff) *)
- := (Const WO~0~0~0~0~0~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).
-Notation "'0x400000000000'" (* 70368744177664 (0x400000000000) *)
- := (Const 70368744177664%Z).
-Notation "'0x400000000000'" (* 70368744177664 (0x400000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000001'" (* 70368744177665 (0x400000000001) *)
- := (Const 70368744177665%Z).
-Notation "'0x400000000001'" (* 70368744177665 (0x400000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffefffffe'" (* 140737471578110 (0x7ffffefffffe) *)
- := (Const 140737471578110%Z).
-Notation "'0x7ffffefffffe'" (* 140737471578110 (0x7ffffefffffe) *)
- := (Const WO~0~0~0~0~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~0~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 "'0x7fffff7fffff'" (* 140737479966719 (0x7fffff7fffff) *)
- := (Const 140737479966719%Z).
-Notation "'0x7fffff7fffff'" (* 140737479966719 (0x7fffff7fffff) *)
- := (Const WO~0~0~0~0~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~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 "'0x7ffffffffdc7'" (* 140737488354759 (0x7ffffffffdc7) *)
- := (Const 140737488354759%Z).
-Notation "'0x7ffffffffdc7'" (* 140737488354759 (0x7ffffffffdc7) *)
- := (Const WO~0~0~0~0~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~0~1~1~1~0~0~0~1~1~1).
-Notation "'0x7ffffffffe1f'" (* 140737488354847 (0x7ffffffffe1f) *)
- := (Const 140737488354847%Z).
-Notation "'0x7ffffffffe1f'" (* 140737488354847 (0x7ffffffffe1f) *)
- := (Const WO~0~0~0~0~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~0~0~0~0~1~1~1~1~1).
-Notation "'0x7fffffffff45'" (* 140737488355141 (0x7fffffffff45) *)
- := (Const 140737488355141%Z).
-Notation "'0x7fffffffff45'" (* 140737488355141 (0x7fffffffff45) *)
- := (Const WO~0~0~0~0~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~0~1~0~0~0~1~0~1).
-Notation "'0x7fffffffffca'" (* 140737488355274 (0x7fffffffffca) *)
- := (Const 140737488355274%Z).
-Notation "'0x7fffffffffca'" (* 140737488355274 (0x7fffffffffca) *)
- := (Const WO~0~0~0~0~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~0~0~1~0~1~0).
-Notation "'0x7fffffffffde'" (* 140737488355294 (0x7fffffffffde) *)
- := (Const 140737488355294%Z).
-Notation "'0x7fffffffffde'" (* 140737488355294 (0x7fffffffffde) *)
- := (Const WO~0~0~0~0~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~0~1~1~1~1~0).
-Notation "'0x7fffffffffe5'" (* 140737488355301 (0x7fffffffffe5) *)
- := (Const 140737488355301%Z).
-Notation "'0x7fffffffffe5'" (* 140737488355301 (0x7fffffffffe5) *)
- := (Const WO~0~0~0~0~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~0~0~1~0~1).
-Notation "'0x7fffffffffe6'" (* 140737488355302 (0x7fffffffffe6) *)
- := (Const 140737488355302%Z).
-Notation "'0x7fffffffffe6'" (* 140737488355302 (0x7fffffffffe6) *)
- := (Const WO~0~0~0~0~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~0~0~1~1~0).
-Notation "'0x7fffffffffe7'" (* 140737488355303 (0x7fffffffffe7) *)
- := (Const 140737488355303%Z).
-Notation "'0x7fffffffffe7'" (* 140737488355303 (0x7fffffffffe7) *)
- := (Const WO~0~0~0~0~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~0~0~1~1~1).
-Notation "'0x7fffffffffee'" (* 140737488355310 (0x7fffffffffee) *)
- := (Const 140737488355310%Z).
-Notation "'0x7fffffffffee'" (* 140737488355310 (0x7fffffffffee) *)
- := (Const WO~0~0~0~0~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~0~1~1~1~0).
-Notation "'0x7ffffffffff1'" (* 140737488355313 (0x7ffffffffff1) *)
- := (Const 140737488355313%Z).
-Notation "'0x7ffffffffff1'" (* 140737488355313 (0x7ffffffffff1) *)
- := (Const WO~0~0~0~0~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~0~0~0~1).
-Notation "'0x7ffffffffff6'" (* 140737488355318 (0x7ffffffffff6) *)
- := (Const 140737488355318%Z).
-Notation "'0x7ffffffffff6'" (* 140737488355318 (0x7ffffffffff6) *)
- := (Const WO~0~0~0~0~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~0~1~1~0).
-Notation "'0x7ffffffffff7'" (* 140737488355319 (0x7ffffffffff7) *)
- := (Const 140737488355319%Z).
-Notation "'0x7ffffffffff7'" (* 140737488355319 (0x7ffffffffff7) *)
- := (Const WO~0~0~0~0~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~0~1~1~1).
-Notation "'0x7ffffffffffe'" (* 140737488355326 (0x7ffffffffffe) *)
- := (Const 140737488355326%Z).
-Notation "'0x7ffffffffffe'" (* 140737488355326 (0x7ffffffffffe) *)
- := (Const WO~0~0~0~0~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~0).
-Notation "'0x7fffffffffff'" (* 140737488355327 (0x7fffffffffff) *)
- := (Const 140737488355327%Z).
-Notation "'0x7fffffffffff'" (* 140737488355327 (0x7fffffffffff) *)
- := (Const WO~0~0~0~0~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).
-Notation "'0x800000000000'" (* 140737488355328 (0x800000000000) *)
- := (Const 140737488355328%Z).
-Notation "'0x800000000000'" (* 140737488355328 (0x800000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000001'" (* 140737488355329 (0x800000000001) *)
- := (Const 140737488355329%Z).
-Notation "'0x800000000001'" (* 140737488355329 (0x800000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xb0ffffffffff'" (* 194613558116351 (0xb0ffffffffff) *)
- := (Const 194613558116351%Z).
-Notation "'0xb0ffffffffff'" (* 194613558116351 (0xb0ffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~1~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).
-Notation "'0xfffaffffffff'" (* 281453501874175 (0xfffaffffffff) *)
- := (Const 281453501874175%Z).
-Notation "'0xfffaffffffff'" (* 281453501874175 (0xfffaffffffff) *)
- := (Const WO~0~0~0~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~0~1~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).
-Notation "'0xfffeffffffff'" (* 281470681743359 (0xfffeffffffff) *)
- := (Const 281470681743359%Z).
-Notation "'0xfffeffffffff'" (* 281470681743359 (0xfffeffffffff) *)
- := (Const WO~0~0~0~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~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).
-Notation "'0xfffffefffffe'" (* 281474959933438 (0xfffffefffffe) *)
- := (Const 281474959933438%Z).
-Notation "'0xfffffefffffe'" (* 281474959933438 (0xfffffefffffe) *)
- := (Const WO~0~0~0~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~0~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 "'0xfffffffeffff'" (* 281474976645119 (0xfffffffeffff) *)
- := (Const 281474976645119%Z).
-Notation "'0xfffffffeffff'" (* 281474976645119 (0xfffffffeffff) *)
- := (Const WO~0~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfffffffffb8e'" (* 281474976709518 (0xfffffffffb8e) *)
- := (Const 281474976709518%Z).
-Notation "'0xfffffffffb8e'" (* 281474976709518 (0xfffffffffb8e) *)
- := (Const WO~0~0~0~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~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0xfffffffffc3e'" (* 281474976709694 (0xfffffffffc3e) *)
- := (Const 281474976709694%Z).
-Notation "'0xfffffffffc3e'" (* 281474976709694 (0xfffffffffc3e) *)
- := (Const WO~0~0~0~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~0~0~0~0~1~1~1~1~1~0).
-Notation "'0xfffffffffe5b'" (* 281474976710235 (0xfffffffffe5b) *)
- := (Const 281474976710235%Z).
-Notation "'0xfffffffffe5b'" (* 281474976710235 (0xfffffffffe5b) *)
- := (Const WO~0~0~0~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~0~0~1~0~1~1~0~1~1).
-Notation "'0xfffffffffe8a'" (* 281474976710282 (0xfffffffffe8a) *)
- := (Const 281474976710282%Z).
-Notation "'0xfffffffffe8a'" (* 281474976710282 (0xfffffffffe8a) *)
- := (Const WO~0~0~0~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~0~1~0~0~0~1~0~1~0).
-Notation "'0xfffffffffec3'" (* 281474976710339 (0xfffffffffec3) *)
- := (Const 281474976710339%Z).
-Notation "'0xfffffffffec3'" (* 281474976710339 (0xfffffffffec3) *)
- := (Const WO~0~0~0~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~0~1~1~0~0~0~0~1~1).
-Notation "'0xffffffffff45'" (* 281474976710469 (0xffffffffff45) *)
- := (Const 281474976710469%Z).
-Notation "'0xffffffffff45'" (* 281474976710469 (0xffffffffff45) *)
- := (Const WO~0~0~0~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~0~1~0~0~0~1~0~1).
-Notation "'0xffffffffff97'" (* 281474976710551 (0xffffffffff97) *)
- := (Const 281474976710551%Z).
-Notation "'0xffffffffff97'" (* 281474976710551 (0xffffffffff97) *)
- := (Const WO~0~0~0~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~0~0~1~0~1~1~1).
-Notation "'0xffffffffffca'" (* 281474976710602 (0xffffffffffca) *)
- := (Const 281474976710602%Z).
-Notation "'0xffffffffffca'" (* 281474976710602 (0xffffffffffca) *)
- := (Const WO~0~0~0~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~0~0~1~0~1~0).
-Notation "'0xffffffffffce'" (* 281474976710606 (0xffffffffffce) *)
- := (Const 281474976710606%Z).
-Notation "'0xffffffffffce'" (* 281474976710606 (0xffffffffffce) *)
- := (Const WO~0~0~0~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~0~0~1~1~1~0).
-Notation "'0xffffffffffe1'" (* 281474976710625 (0xffffffffffe1) *)
- := (Const 281474976710625%Z).
-Notation "'0xffffffffffe1'" (* 281474976710625 (0xffffffffffe1) *)
- := (Const WO~0~0~0~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~0~0~0~0~1).
-Notation "'0xffffffffffe2'" (* 281474976710626 (0xffffffffffe2) *)
- := (Const 281474976710626%Z).
-Notation "'0xffffffffffe2'" (* 281474976710626 (0xffffffffffe2) *)
- := (Const WO~0~0~0~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~0~0~0~1~0).
-Notation "'0xffffffffffe7'" (* 281474976710631 (0xffffffffffe7) *)
- := (Const 281474976710631%Z).
-Notation "'0xffffffffffe7'" (* 281474976710631 (0xffffffffffe7) *)
- := (Const WO~0~0~0~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~0~0~1~1~1).
-Notation "'0xffffffffffed'" (* 281474976710637 (0xffffffffffed) *)
- := (Const 281474976710637%Z).
-Notation "'0xffffffffffed'" (* 281474976710637 (0xffffffffffed) *)
- := (Const WO~0~0~0~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~0~1~1~0~1).
-Notation "'0xffffffffffee'" (* 281474976710638 (0xffffffffffee) *)
- := (Const 281474976710638%Z).
-Notation "'0xffffffffffee'" (* 281474976710638 (0xffffffffffee) *)
- := (Const WO~0~0~0~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~0~1~1~1~0).
-Notation "'0xffffffffffef'" (* 281474976710639 (0xffffffffffef) *)
- := (Const 281474976710639%Z).
-Notation "'0xffffffffffef'" (* 281474976710639 (0xffffffffffef) *)
- := (Const WO~0~0~0~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~0~1~1~1~1).
-Notation "'0xfffffffffff5'" (* 281474976710645 (0xfffffffffff5) *)
- := (Const 281474976710645%Z).
-Notation "'0xfffffffffff5'" (* 281474976710645 (0xfffffffffff5) *)
- := (Const WO~0~0~0~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~0~1~0~1).
-Notation "'0xfffffffffff7'" (* 281474976710647 (0xfffffffffff7) *)
- := (Const 281474976710647%Z).
-Notation "'0xfffffffffff7'" (* 281474976710647 (0xfffffffffff7) *)
- := (Const WO~0~0~0~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~0~1~1~1).
-Notation "'0xfffffffffffd'" (* 281474976710653 (0xfffffffffffd) *)
- := (Const 281474976710653%Z).
-Notation "'0xfffffffffffd'" (* 281474976710653 (0xfffffffffffd) *)
- := (Const WO~0~0~0~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~0~1).
-Notation "'0xfffffffffffe'" (* 281474976710654 (0xfffffffffffe) *)
- := (Const 281474976710654%Z).
-Notation "'0xfffffffffffe'" (* 281474976710654 (0xfffffffffffe) *)
- := (Const WO~0~0~0~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~0).
-Notation "'0xffffffffffff'" (* 281474976710655 (0xffffffffffff) *)
- := (Const 281474976710655%Z).
-Notation "'0xffffffffffff'" (* 281474976710655 (0xffffffffffff) *)
- := (Const WO~0~0~0~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).
-Notation "'0x1000000000000'" (* 281474976710656 (0x1000000000000) *)
- := (Const 281474976710656%Z).
-Notation "'0x1000000000000'" (* 281474976710656 (0x1000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000001'" (* 281474976710657 (0x1000000000001) *)
- := (Const 281474976710657%Z).
-Notation "'0x1000000000001'" (* 281474976710657 (0x1000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x161fffffffffe'" (* 389227116232702 (0x161fffffffffe) *)
- := (Const 389227116232702%Z).
-Notation "'0x161fffffffffe'" (* 389227116232702 (0x161fffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~1~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~0).
-Notation "'0x1fff5fffffffe'" (* 562907003748350 (0x1fff5fffffffe) *)
- := (Const 562907003748350%Z).
-Notation "'0x1fff5fffffffe'" (* 562907003748350 (0x1fff5fffffffe) *)
- := (Const WO~0~0~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~0~1~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~0).
-Notation "'0x1fffffffdfffe'" (* 562949953290238 (0x1fffffffdfffe) *)
- := (Const 562949953290238%Z).
-Notation "'0x1fffffffdfffe'" (* 562949953290238 (0x1fffffffdfffe) *)
- := (Const WO~0~0~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0).
-Notation "'0x1fffffffffcb6'" (* 562949953420470 (0x1fffffffffcb6) *)
- := (Const 562949953420470%Z).
-Notation "'0x1fffffffffcb6'" (* 562949953420470 (0x1fffffffffcb6) *)
- := (Const WO~0~0~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~0~0~1~0~1~1~0~1~1~0).
-Notation "'0x1fffffffffd86'" (* 562949953420678 (0x1fffffffffd86) *)
- := (Const 562949953420678%Z).
-Notation "'0x1fffffffffd86'" (* 562949953420678 (0x1fffffffffd86) *)
- := (Const WO~0~0~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~0~1~1~0~0~0~0~1~1~0).
-Notation "'0x1fffffffffe8a'" (* 562949953420938 (0x1fffffffffe8a) *)
- := (Const 562949953420938%Z).
-Notation "'0x1fffffffffe8a'" (* 562949953420938 (0x1fffffffffe8a) *)
- := (Const WO~0~0~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~0~1~0~0~0~1~0~1~0).
-Notation "'0x1ffffffffff2e'" (* 562949953421102 (0x1ffffffffff2e) *)
- := (Const 562949953421102%Z).
-Notation "'0x1ffffffffff2e'" (* 562949953421102 (0x1ffffffffff2e) *)
- := (Const WO~0~0~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~0~0~1~0~1~1~1~0).
-Notation "'0x1ffffffffffc2'" (* 562949953421250 (0x1ffffffffffc2) *)
- := (Const 562949953421250%Z).
-Notation "'0x1ffffffffffc2'" (* 562949953421250 (0x1ffffffffffc2) *)
- := (Const WO~0~0~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~0~0~0~0~1~0).
-Notation "'0x1ffffffffffce'" (* 562949953421262 (0x1ffffffffffce) *)
- := (Const 562949953421262%Z).
-Notation "'0x1ffffffffffce'" (* 562949953421262 (0x1ffffffffffce) *)
- := (Const WO~0~0~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~0~0~1~1~1~0).
-Notation "'0x1ffffffffffda'" (* 562949953421274 (0x1ffffffffffda) *)
- := (Const 562949953421274%Z).
-Notation "'0x1ffffffffffda'" (* 562949953421274 (0x1ffffffffffda) *)
- := (Const WO~0~0~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~0~1~1~0~1~0).
-Notation "'0x1ffffffffffde'" (* 562949953421278 (0x1ffffffffffde) *)
- := (Const 562949953421278%Z).
-Notation "'0x1ffffffffffde'" (* 562949953421278 (0x1ffffffffffde) *)
- := (Const WO~0~0~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~0~1~1~1~1~0).
-Notation "'0x1ffffffffffdf'" (* 562949953421279 (0x1ffffffffffdf) *)
- := (Const 562949953421279%Z).
-Notation "'0x1ffffffffffdf'" (* 562949953421279 (0x1ffffffffffdf) *)
- := (Const WO~0~0~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~0~1~1~1~1~1).
-Notation "'0x1ffffffffffea'" (* 562949953421290 (0x1ffffffffffea) *)
- := (Const 562949953421290%Z).
-Notation "'0x1ffffffffffea'" (* 562949953421290 (0x1ffffffffffea) *)
- := (Const WO~0~0~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~0~1~0~1~0).
-Notation "'0x1ffffffffffeb'" (* 562949953421291 (0x1ffffffffffeb) *)
- := (Const 562949953421291%Z).
-Notation "'0x1ffffffffffeb'" (* 562949953421291 (0x1ffffffffffeb) *)
- := (Const WO~0~0~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~0~1~0~1~1).
-Notation "'0x1ffffffffffed'" (* 562949953421293 (0x1ffffffffffed) *)
- := (Const 562949953421293%Z).
-Notation "'0x1ffffffffffed'" (* 562949953421293 (0x1ffffffffffed) *)
- := (Const WO~0~0~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~0~1~1~0~1).
-Notation "'0x1ffffffffffee'" (* 562949953421294 (0x1ffffffffffee) *)
- := (Const 562949953421294%Z).
-Notation "'0x1ffffffffffee'" (* 562949953421294 (0x1ffffffffffee) *)
- := (Const WO~0~0~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~0~1~1~1~0).
-Notation "'0x1ffffffffffef'" (* 562949953421295 (0x1ffffffffffef) *)
- := (Const 562949953421295%Z).
-Notation "'0x1ffffffffffef'" (* 562949953421295 (0x1ffffffffffef) *)
- := (Const WO~0~0~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~0~1~1~1~1).
-Notation "'0x1fffffffffff1'" (* 562949953421297 (0x1fffffffffff1) *)
- := (Const 562949953421297%Z).
-Notation "'0x1fffffffffff1'" (* 562949953421297 (0x1fffffffffff1) *)
- := (Const WO~0~0~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~0~0~0~1).
-Notation "'0x1fffffffffff7'" (* 562949953421303 (0x1fffffffffff7) *)
- := (Const 562949953421303%Z).
-Notation "'0x1fffffffffff7'" (* 562949953421303 (0x1fffffffffff7) *)
- := (Const WO~0~0~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~0~1~1~1).
-Notation "'0x1fffffffffffa'" (* 562949953421306 (0x1fffffffffffa) *)
- := (Const 562949953421306%Z).
-Notation "'0x1fffffffffffa'" (* 562949953421306 (0x1fffffffffffa) *)
- := (Const WO~0~0~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~0~1~0).
-Notation "'0x1fffffffffffc'" (* 562949953421308 (0x1fffffffffffc) *)
- := (Const 562949953421308%Z).
-Notation "'0x1fffffffffffc'" (* 562949953421308 (0x1fffffffffffc) *)
- := (Const WO~0~0~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~0~0).
-Notation "'0x1fffffffffffe'" (* 562949953421310 (0x1fffffffffffe) *)
- := (Const 562949953421310%Z).
-Notation "'0x1fffffffffffe'" (* 562949953421310 (0x1fffffffffffe) *)
- := (Const WO~0~0~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~0).
-Notation "'0x1ffffffffffff'" (* 562949953421311 (0x1ffffffffffff) *)
- := (Const 562949953421311%Z).
-Notation "'0x1ffffffffffff'" (* 562949953421311 (0x1ffffffffffff) *)
- := (Const WO~0~0~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).
-Notation "'0x2000000000000'" (* 562949953421312 (0x2000000000000) *)
- := (Const 562949953421312%Z).
-Notation "'0x2000000000000'" (* 562949953421312 (0x2000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000001'" (* 562949953421313 (0x2000000000001) *)
- := (Const 562949953421313%Z).
-Notation "'0x2000000000001'" (* 562949953421313 (0x2000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffdffffff'" (* 1125899873288191 (0x3fffffdffffff) *)
- := (Const 1125899873288191%Z).
-Notation "'0x3fffffdffffff'" (* 1125899873288191 (0x3fffffdffffff) *)
- := (Const WO~0~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~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 "'0x3ffffffffffbe'" (* 1125899906842558 (0x3ffffffffffbe) *)
- := (Const 1125899906842558%Z).
-Notation "'0x3ffffffffffbe'" (* 1125899906842558 (0x3ffffffffffbe) *)
- := (Const WO~0~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~0~1~1~1~1~1~0).
-Notation "'0x3ffffffffffd6'" (* 1125899906842582 (0x3ffffffffffd6) *)
- := (Const 1125899906842582%Z).
-Notation "'0x3ffffffffffd6'" (* 1125899906842582 (0x3ffffffffffd6) *)
- := (Const WO~0~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~0~1~0~1~1~0).
-Notation "'0x3ffffffffffda'" (* 1125899906842586 (0x3ffffffffffda) *)
- := (Const 1125899906842586%Z).
-Notation "'0x3ffffffffffda'" (* 1125899906842586 (0x3ffffffffffda) *)
- := (Const WO~0~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~0~1~1~0~1~0).
-Notation "'0x3ffffffffffde'" (* 1125899906842590 (0x3ffffffffffde) *)
- := (Const 1125899906842590%Z).
-Notation "'0x3ffffffffffde'" (* 1125899906842590 (0x3ffffffffffde) *)
- := (Const WO~0~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~0~1~1~1~1~0).
-Notation "'0x3ffffffffffe1'" (* 1125899906842593 (0x3ffffffffffe1) *)
- := (Const 1125899906842593%Z).
-Notation "'0x3ffffffffffe1'" (* 1125899906842593 (0x3ffffffffffe1) *)
- := (Const WO~0~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~0~0~0~0~1).
-Notation "'0x3ffffffffffe2'" (* 1125899906842594 (0x3ffffffffffe2) *)
- := (Const 1125899906842594%Z).
-Notation "'0x3ffffffffffe2'" (* 1125899906842594 (0x3ffffffffffe2) *)
- := (Const WO~0~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~0~0~0~1~0).
-Notation "'0x3ffffffffffee'" (* 1125899906842606 (0x3ffffffffffee) *)
- := (Const 1125899906842606%Z).
-Notation "'0x3ffffffffffee'" (* 1125899906842606 (0x3ffffffffffee) *)
- := (Const WO~0~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~0~1~1~1~0).
-Notation "'0x3ffffffffffef'" (* 1125899906842607 (0x3ffffffffffef) *)
- := (Const 1125899906842607%Z).
-Notation "'0x3ffffffffffef'" (* 1125899906842607 (0x3ffffffffffef) *)
- := (Const WO~0~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~0~1~1~1~1).
-Notation "'0x3fffffffffffb'" (* 1125899906842619 (0x3fffffffffffb) *)
- := (Const 1125899906842619%Z).
-Notation "'0x3fffffffffffb'" (* 1125899906842619 (0x3fffffffffffb) *)
- := (Const WO~0~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~0~1~1).
-Notation "'0x3fffffffffffd'" (* 1125899906842621 (0x3fffffffffffd) *)
- := (Const 1125899906842621%Z).
-Notation "'0x3fffffffffffd'" (* 1125899906842621 (0x3fffffffffffd) *)
- := (Const WO~0~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~0~1).
-Notation "'0x3fffffffffffe'" (* 1125899906842622 (0x3fffffffffffe) *)
- := (Const 1125899906842622%Z).
-Notation "'0x3fffffffffffe'" (* 1125899906842622 (0x3fffffffffffe) *)
- := (Const WO~0~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~0).
-Notation "'0x3ffffffffffff'" (* 1125899906842623 (0x3ffffffffffff) *)
- := (Const 1125899906842623%Z).
-Notation "'0x3ffffffffffff'" (* 1125899906842623 (0x3ffffffffffff) *)
- := (Const WO~0~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).
-Notation "'0x4000000000000'" (* 1125899906842624 (0x4000000000000) *)
- := (Const 1125899906842624%Z).
-Notation "'0x4000000000000'" (* 1125899906842624 (0x4000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000001'" (* 1125899906842625 (0x4000000000001) *)
- := (Const 1125899906842625%Z).
-Notation "'0x4000000000001'" (* 1125899906842625 (0x4000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x52fffffffffff'" (* 1460151441686527 (0x52fffffffffff) *)
- := (Const 1460151441686527%Z).
-Notation "'0x52fffffffffff'" (* 1460151441686527 (0x52fffffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~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).
-Notation "'0x7dbbfffffffff'" (* 2211942517178367 (0x7dbbfffffffff) *)
- := (Const 2211942517178367%Z).
-Notation "'0x7dbbfffffffff'" (* 2211942517178367 (0x7dbbfffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~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).
-Notation "'0x7f0a7ffffffff'" (* 2234929182146559 (0x7f0a7ffffffff) *)
- := (Const 2234929182146559%Z).
-Notation "'0x7f0a7ffffffff'" (* 2234929182146559 (0x7f0a7ffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~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).
-Notation "'0x7fd3fffffffff'" (* 2248776156708863 (0x7fd3fffffffff) *)
- := (Const 2248776156708863%Z).
-Notation "'0x7fd3fffffffff'" (* 2248776156708863 (0x7fd3fffffffff) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~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).
-Notation "'0x7fffffff80000'" (* 2251799813160960 (0x7fffffff80000) *)
- := (Const 2251799813160960%Z).
-Notation "'0x7fffffff80000'" (* 2251799813160960 (0x7fffffff80000) *)
- := (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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x7fffffffffd03'" (* 2251799813684483 (0x7fffffffffd03) *)
- := (Const 2251799813684483%Z).
-Notation "'0x7fffffffffd03'" (* 2251799813684483 (0x7fffffffffd03) *)
- := (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~0~1~0~0~0~0~0~0~1~1).
-Notation "'0x7ffffffffffc2'" (* 2251799813685186 (0x7ffffffffffc2) *)
- := (Const 2251799813685186%Z).
-Notation "'0x7ffffffffffc2'" (* 2251799813685186 (0x7ffffffffffc2) *)
- := (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~0~0~0~0~1~0).
-Notation "'0x7ffffffffffda'" (* 2251799813685210 (0x7ffffffffffda) *)
- := (Const 2251799813685210%Z).
-Notation "'0x7ffffffffffda'" (* 2251799813685210 (0x7ffffffffffda) *)
- := (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~0~1~1~0~1~0).
-Notation "'0x7ffffffffffde'" (* 2251799813685214 (0x7ffffffffffde) *)
- := (Const 2251799813685214%Z).
-Notation "'0x7ffffffffffde'" (* 2251799813685214 (0x7ffffffffffde) *)
- := (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~0~1~1~1~1~0).
-Notation "'0x7ffffffffffe1'" (* 2251799813685217 (0x7ffffffffffe1) *)
- := (Const 2251799813685217%Z).
-Notation "'0x7ffffffffffe1'" (* 2251799813685217 (0x7ffffffffffe1) *)
- := (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~0~0~0~0~1).
-Notation "'0x7ffffffffffed'" (* 2251799813685229 (0x7ffffffffffed) *)
- := (Const 2251799813685229%Z).
-Notation "'0x7ffffffffffed'" (* 2251799813685229 (0x7ffffffffffed) *)
- := (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~0~1~1~0~1).
-Notation "'0x7ffffffffffef'" (* 2251799813685231 (0x7ffffffffffef) *)
- := (Const 2251799813685231%Z).
-Notation "'0x7ffffffffffef'" (* 2251799813685231 (0x7ffffffffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0x7fffffffffff6'" (* 2251799813685238 (0x7fffffffffff6) *)
- := (Const 2251799813685238%Z).
-Notation "'0x7fffffffffff6'" (* 2251799813685238 (0x7fffffffffff6) *)
- := (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~0~1~1~0).
-Notation "'0x7fffffffffff7'" (* 2251799813685239 (0x7fffffffffff7) *)
- := (Const 2251799813685239%Z).
-Notation "'0x7fffffffffff7'" (* 2251799813685239 (0x7fffffffffff7) *)
- := (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~0~1~1~1).
-Notation "'0x7fffffffffffa'" (* 2251799813685242 (0x7fffffffffffa) *)
- := (Const 2251799813685242%Z).
-Notation "'0x7fffffffffffa'" (* 2251799813685242 (0x7fffffffffffa) *)
- := (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~0~1~0).
-Notation "'0x7fffffffffffd'" (* 2251799813685245 (0x7fffffffffffd) *)
- := (Const 2251799813685245%Z).
-Notation "'0x7fffffffffffd'" (* 2251799813685245 (0x7fffffffffffd) *)
- := (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~0~1).
-Notation "'0x7fffffffffffe'" (* 2251799813685246 (0x7fffffffffffe) *)
- := (Const 2251799813685246%Z).
-Notation "'0x7fffffffffffe'" (* 2251799813685246 (0x7fffffffffffe) *)
- := (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~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 "'0x8000000000000'" (* 2251799813685248 (0x8000000000000) *)
- := (Const 2251799813685248%Z).
-Notation "'0x8000000000000'" (* 2251799813685248 (0x8000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000001'" (* 2251799813685249 (0x8000000000001) *)
- := (Const 2251799813685249%Z).
-Notation "'0x8000000000001'" (* 2251799813685249 (0x8000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xa5ffffffffffe'" (* 2920302883373054 (0xa5ffffffffffe) *)
- := (Const 2920302883373054%Z).
-Notation "'0xa5ffffffffffe'" (* 2920302883373054 (0xa5ffffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~1~0~0~1~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~0).
-Notation "'0xfb77ffffffffe'" (* 4423885034356734 (0xfb77ffffffffe) *)
- := (Const 4423885034356734%Z).
-Notation "'0xfb77ffffffffe'" (* 4423885034356734 (0xfb77ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~0~1~1~0~1~1~1~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~0).
-Notation "'0xfe14ffffffffe'" (* 4469858364293118 (0xfe14ffffffffe) *)
- := (Const 4469858364293118%Z).
-Notation "'0xfe14ffffffffe'" (* 4469858364293118 (0xfe14ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~0~0~0~0~1~0~1~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~0).
-Notation "'0xffa7ffffffffe'" (* 4497552313417726 (0xffa7ffffffffe) *)
- := (Const 4497552313417726%Z).
-Notation "'0xffa7ffffffffe'" (* 4497552313417726 (0xffa7ffffffffe) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~0~1~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~0).
-Notation "'0xffffefffffc2f'" (* 4503595332402223 (0xffffefffffc2f) *)
- := (Const 4503595332402223%Z).
-Notation "'0xffffefffffc2f'" (* 4503595332402223 (0xffffefffffc2f) *)
- := (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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0xffffffffffa06'" (* 4503599627368966 (0xffffffffffa06) *)
- := (Const 4503599627368966%Z).
-Notation "'0xffffffffffa06'" (* 4503599627368966 (0xffffffffffa06) *)
- := (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~0~1~0~0~0~0~0~0~1~1~0).
-Notation "'0xffffffffffdc7'" (* 4503599627369927 (0xffffffffffdc7) *)
- := (Const 4503599627369927%Z).
-Notation "'0xffffffffffdc7'" (* 4503599627369927 (0xffffffffffdc7) *)
- := (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~0~1~1~1~0~0~0~1~1~1).
-Notation "'0xffffffffffe1f'" (* 4503599627370015 (0xffffffffffe1f) *)
- := (Const 4503599627370015%Z).
-Notation "'0xffffffffffe1f'" (* 4503599627370015 (0xffffffffffe1f) *)
- := (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~0~0~0~0~1~1~1~1~1).
-Notation "'0xfffffffffff43'" (* 4503599627370307 (0xfffffffffff43) *)
- := (Const 4503599627370307%Z).
-Notation "'0xfffffffffff43'" (* 4503599627370307 (0xfffffffffff43) *)
- := (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~0~1~0~0~0~0~1~1).
-Notation "'0xfffffffffff45'" (* 4503599627370309 (0xfffffffffff45) *)
- := (Const 4503599627370309%Z).
-Notation "'0xfffffffffff45'" (* 4503599627370309 (0xfffffffffff45) *)
- := (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~0~1~0~0~0~1~0~1).
-Notation "'0xfffffffffffc2'" (* 4503599627370434 (0xfffffffffffc2) *)
- := (Const 4503599627370434%Z).
-Notation "'0xfffffffffffc2'" (* 4503599627370434 (0xfffffffffffc2) *)
- := (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~0~0~0~1~0).
-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 "'0xfffffffffffde'" (* 4503599627370462 (0xfffffffffffde) *)
- := (Const 4503599627370462%Z).
-Notation "'0xfffffffffffde'" (* 4503599627370462 (0xfffffffffffde) *)
- := (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~1~1~0).
-Notation "'0xfffffffffffeb'" (* 4503599627370475 (0xfffffffffffeb) *)
- := (Const 4503599627370475%Z).
-Notation "'0xfffffffffffeb'" (* 4503599627370475 (0xfffffffffffeb) *)
- := (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~0~1~0~1~1).
-Notation "'0xfffffffffffee'" (* 4503599627370478 (0xfffffffffffee) *)
- := (Const 4503599627370478%Z).
-Notation "'0xfffffffffffee'" (* 4503599627370478 (0xfffffffffffee) *)
- := (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~0~1~1~1~0).
-Notation "'0xfffffffffffef'" (* 4503599627370479 (0xfffffffffffef) *)
- := (Const 4503599627370479%Z).
-Notation "'0xfffffffffffef'" (* 4503599627370479 (0xfffffffffffef) *)
- := (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~0~1~1~1~1).
-Notation "'0xffffffffffffa'" (* 4503599627370490 (0xffffffffffffa) *)
- := (Const 4503599627370490%Z).
-Notation "'0xffffffffffffa'" (* 4503599627370490 (0xffffffffffffa) *)
- := (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~0~1~0).
-Notation "'0xffffffffffffb'" (* 4503599627370491 (0xffffffffffffb) *)
- := (Const 4503599627370491%Z).
-Notation "'0xffffffffffffb'" (* 4503599627370491 (0xffffffffffffb) *)
- := (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~0~1~1).
-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).
-Notation "'0xfffffffffffff'" (* 4503599627370495 (0xfffffffffffff) *)
- := (Const 4503599627370495%Z).
-Notation "'0xfffffffffffff'" (* 4503599627370495 (0xfffffffffffff) *)
- := (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~1).
-Notation "'0x10000000000000'" (* 4503599627370496 (0x10000000000000) *)
- := (Const 4503599627370496%Z).
-Notation "'0x10000000000000'" (* 4503599627370496 (0x10000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000001'" (* 4503599627370497 (0x10000000000001) *)
- := (Const 4503599627370497%Z).
-Notation "'0x10000000000001'" (* 4503599627370497 (0x10000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffdfffff85e'" (* 9007190664804446 (0x1ffffdfffff85e) *)
- := (Const 9007190664804446%Z).
-Notation "'0x1ffffdfffff85e'" (* 9007190664804446 (0x1ffffdfffff85e) *)
- := (Const WO~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~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1~0).
-Notation "'0x1ffffffbffffff'" (* 9007199187632127 (0x1ffffffbffffff) *)
- := (Const 9007199187632127%Z).
-Notation "'0x1ffffffbffffff'" (* 9007199187632127 (0x1ffffffbffffff) *)
- := (Const WO~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~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 "'0x1ffffffffffb8e'" (* 9007199254739854 (0x1ffffffffffb8e) *)
- := (Const 9007199254739854%Z).
-Notation "'0x1ffffffffffb8e'" (* 9007199254739854 (0x1ffffffffffb8e) *)
- := (Const WO~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~0~1~1~1~0~0~0~1~1~1~0).
-Notation "'0x1ffffffffffc3e'" (* 9007199254740030 (0x1ffffffffffc3e) *)
- := (Const 9007199254740030%Z).
-Notation "'0x1ffffffffffc3e'" (* 9007199254740030 (0x1ffffffffffc3e) *)
- := (Const WO~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~0~0~0~0~1~1~1~1~1~0).
-Notation "'0x1ffffffffffe86'" (* 9007199254740614 (0x1ffffffffffe86) *)
- := (Const 9007199254740614%Z).
-Notation "'0x1ffffffffffe86'" (* 9007199254740614 (0x1ffffffffffe86) *)
- := (Const WO~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~0~1~0~0~0~0~1~1~0).
-Notation "'0x1ffffffffffe8a'" (* 9007199254740618 (0x1ffffffffffe8a) *)
- := (Const 9007199254740618%Z).
-Notation "'0x1ffffffffffe8a'" (* 9007199254740618 (0x1ffffffffffe8a) *)
- := (Const WO~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~0~1~0~0~0~1~0~1~0).
-Notation "'0x1fffffffffffd6'" (* 9007199254740950 (0x1fffffffffffd6) *)
- := (Const 9007199254740950%Z).
-Notation "'0x1fffffffffffd6'" (* 9007199254740950 (0x1fffffffffffd6) *)
- := (Const WO~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~0~1~0~1~1~0).
-Notation "'0x1fffffffffffde'" (* 9007199254740958 (0x1fffffffffffde) *)
- := (Const 9007199254740958%Z).
-Notation "'0x1fffffffffffde'" (* 9007199254740958 (0x1fffffffffffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0x1fffffffffffe3'" (* 9007199254740963 (0x1fffffffffffe3) *)
- := (Const 9007199254740963%Z).
-Notation "'0x1fffffffffffe3'" (* 9007199254740963 (0x1fffffffffffe3) *)
- := (Const WO~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~0~0~0~1~1).
-Notation "'0x1fffffffffffe7'" (* 9007199254740967 (0x1fffffffffffe7) *)
- := (Const 9007199254740967%Z).
-Notation "'0x1fffffffffffe7'" (* 9007199254740967 (0x1fffffffffffe7) *)
- := (Const WO~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~0~0~1~1~1).
-Notation "'0x1ffffffffffff1'" (* 9007199254740977 (0x1ffffffffffff1) *)
- := (Const 9007199254740977%Z).
-Notation "'0x1ffffffffffff1'" (* 9007199254740977 (0x1ffffffffffff1) *)
- := (Const WO~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~0~0~0~1).
-Notation "'0x1ffffffffffff6'" (* 9007199254740982 (0x1ffffffffffff6) *)
- := (Const 9007199254740982%Z).
-Notation "'0x1ffffffffffff6'" (* 9007199254740982 (0x1ffffffffffff6) *)
- := (Const WO~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~0~1~1~0).
-Notation "'0x1ffffffffffffc'" (* 9007199254740988 (0x1ffffffffffffc) *)
- := (Const 9007199254740988%Z).
-Notation "'0x1ffffffffffffc'" (* 9007199254740988 (0x1ffffffffffffc) *)
- := (Const WO~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~0).
-Notation "'0x1ffffffffffffe'" (* 9007199254740990 (0x1ffffffffffffe) *)
- := (Const 9007199254740990%Z).
-Notation "'0x1ffffffffffffe'" (* 9007199254740990 (0x1ffffffffffffe) *)
- := (Const WO~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~1~0).
-Notation "'0x1fffffffffffff'" (* 9007199254740991 (0x1fffffffffffff) *)
- := (Const 9007199254740991%Z).
-Notation "'0x1fffffffffffff'" (* 9007199254740991 (0x1fffffffffffff) *)
- := (Const WO~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~1~1).
-Notation "'0x20000000000000'" (* 9007199254740992 (0x20000000000000) *)
- := (Const 9007199254740992%Z).
-Notation "'0x20000000000000'" (* 9007199254740992 (0x20000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000001'" (* 9007199254740993 (0x20000000000001) *)
- := (Const 9007199254740993%Z).
-Notation "'0x20000000000001'" (* 9007199254740993 (0x20000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffc6'" (* 18014398509481926 (0x3fffffffffffc6) *)
- := (Const 18014398509481926%Z).
-Notation "'0x3fffffffffffc6'" (* 18014398509481926 (0x3fffffffffffc6) *)
- := (Const WO~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~0~0~0~1~1~0).
-Notation "'0x3fffffffffffce'" (* 18014398509481934 (0x3fffffffffffce) *)
- := (Const 18014398509481934%Z).
-Notation "'0x3fffffffffffce'" (* 18014398509481934 (0x3fffffffffffce) *)
- := (Const WO~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~0~0~1~1~1~0).
-Notation "'0x3fffffffffffe2'" (* 18014398509481954 (0x3fffffffffffe2) *)
- := (Const 18014398509481954%Z).
-Notation "'0x3fffffffffffe2'" (* 18014398509481954 (0x3fffffffffffe2) *)
- := (Const WO~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~0~0~0~1~0).
-Notation "'0x3ffffffffffff7'" (* 18014398509481975 (0x3ffffffffffff7) *)
- := (Const 18014398509481975%Z).
-Notation "'0x3ffffffffffff7'" (* 18014398509481975 (0x3ffffffffffff7) *)
- := (Const WO~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~0~1~1~1).
-Notation "'0x3ffffffffffffd'" (* 18014398509481981 (0x3ffffffffffffd) *)
- := (Const 18014398509481981%Z).
-Notation "'0x3ffffffffffffd'" (* 18014398509481981 (0x3ffffffffffffd) *)
- := (Const WO~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~1~0~1).
-Notation "'0x3ffffffffffffe'" (* 18014398509481982 (0x3ffffffffffffe) *)
- := (Const 18014398509481982%Z).
-Notation "'0x3ffffffffffffe'" (* 18014398509481982 (0x3ffffffffffffe) *)
- := (Const WO~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~1~1~0).
-Notation "'0x3fffffffffffff'" (* 18014398509481983 (0x3fffffffffffff) *)
- := (Const 18014398509481983%Z).
-Notation "'0x3fffffffffffff'" (* 18014398509481983 (0x3fffffffffffff) *)
- := (Const WO~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~1~1~1).
-Notation "'0x40000000000000'" (* 18014398509481984 (0x40000000000000) *)
- := (Const 18014398509481984%Z).
-Notation "'0x40000000000000'" (* 18014398509481984 (0x40000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000001'" (* 18014398509481985 (0x40000000000001) *)
- := (Const 18014398509481985%Z).
-Notation "'0x40000000000001'" (* 18014398509481985 (0x40000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffe5b'" (* 36028797018963547 (0x7ffffffffffe5b) *)
- := (Const 36028797018963547%Z).
-Notation "'0x7ffffffffffe5b'" (* 36028797018963547 (0x7ffffffffffe5b) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~1~0~1~1~0~1~1).
-Notation "'0x7ffffffffffec3'" (* 36028797018963651 (0x7ffffffffffec3) *)
- := (Const 36028797018963651%Z).
-Notation "'0x7ffffffffffec3'" (* 36028797018963651 (0x7ffffffffffec3) *)
- := (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~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~0~0~0~1~1).
-Notation "'0x7fffffffffff45'" (* 36028797018963781 (0x7fffffffffff45) *)
- := (Const 36028797018963781%Z).
-Notation "'0x7fffffffffff45'" (* 36028797018963781 (0x7fffffffffff45) *)
- := (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~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~0~0~1~0~1).
-Notation "'0x7fffffffffff97'" (* 36028797018963863 (0x7fffffffffff97) *)
- := (Const 36028797018963863%Z).
-Notation "'0x7fffffffffff97'" (* 36028797018963863 (0x7fffffffffff97) *)
- := (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~1~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~0~1~0~1~1~1).
-Notation "'0x7fffffffffffe1'" (* 36028797018963937 (0x7fffffffffffe1) *)
- := (Const 36028797018963937%Z).
-Notation "'0x7fffffffffffe1'" (* 36028797018963937 (0x7fffffffffffe1) *)
- := (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~1~1~1~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~0~0~0~1).
-Notation "'0x7fffffffffffe7'" (* 36028797018963943 (0x7fffffffffffe7) *)
- := (Const 36028797018963943%Z).
-Notation "'0x7fffffffffffe7'" (* 36028797018963943 (0x7fffffffffffe7) *)
- := (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~1~1~1~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~0~1~1~1).
-Notation "'0x7fffffffffffeb'" (* 36028797018963947 (0x7fffffffffffeb) *)
- := (Const 36028797018963947%Z).
-Notation "'0x7fffffffffffeb'" (* 36028797018963947 (0x7fffffffffffeb) *)
- := (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~1~1~1~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~1~1).
-Notation "'0x7fffffffffffed'" (* 36028797018963949 (0x7fffffffffffed) *)
- := (Const 36028797018963949%Z).
-Notation "'0x7fffffffffffed'" (* 36028797018963949 (0x7fffffffffffed) *)
- := (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~1~1~1~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).
-Notation "'0x7fffffffffffee'" (* 36028797018963950 (0x7fffffffffffee) *)
- := (Const 36028797018963950%Z).
-Notation "'0x7fffffffffffee'" (* 36028797018963950 (0x7fffffffffffee) *)
- := (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~1~1~1~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~1~0).
-Notation "'0x7fffffffffffef'" (* 36028797018963951 (0x7fffffffffffef) *)
- := (Const 36028797018963951%Z).
-Notation "'0x7fffffffffffef'" (* 36028797018963951 (0x7fffffffffffef) *)
- := (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~1~1~1~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~1~1).
-Notation "'0x7ffffffffffffa'" (* 36028797018963962 (0x7ffffffffffffa) *)
- := (Const 36028797018963962%Z).
-Notation "'0x7ffffffffffffa'" (* 36028797018963962 (0x7ffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0x7ffffffffffffc'" (* 36028797018963964 (0x7ffffffffffffc) *)
- := (Const 36028797018963964%Z).
-Notation "'0x7ffffffffffffc'" (* 36028797018963964 (0x7ffffffffffffc) *)
- := (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~1~1~1~1~1~1~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~0).
-Notation "'0x7ffffffffffffe'" (* 36028797018963966 (0x7ffffffffffffe) *)
- := (Const 36028797018963966%Z).
-Notation "'0x7ffffffffffffe'" (* 36028797018963966 (0x7ffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0x7fffffffffffff'" (* 36028797018963967 (0x7fffffffffffff) *)
- := (Const 36028797018963967%Z).
-Notation "'0x7fffffffffffff'" (* 36028797018963967 (0x7fffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0x80000000000000'" (* 36028797018963968 (0x80000000000000) *)
- := (Const 36028797018963968%Z).
-Notation "'0x80000000000000'" (* 36028797018963968 (0x80000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000001'" (* 36028797018963969 (0x80000000000001) *)
- := (Const 36028797018963969%Z).
-Notation "'0x80000000000001'" (* 36028797018963969 (0x80000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffff0000000000'" (* 72056494526300160 (0xffff0000000000) *)
- := (Const 72056494526300160%Z).
-Notation "'0xffff0000000000'" (* 72056494526300160 (0xffff0000000000) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0xfffffffffffcb6'" (* 72057594037927094 (0xfffffffffffcb6) *)
- := (Const 72057594037927094%Z).
-Notation "'0xfffffffffffcb6'" (* 72057594037927094 (0xfffffffffffcb6) *)
- := (Const WO~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~0~1~0~1~1~0~1~1~0).
-Notation "'0xfffffffffffd86'" (* 72057594037927302 (0xfffffffffffd86) *)
- := (Const 72057594037927302%Z).
-Notation "'0xfffffffffffd86'" (* 72057594037927302 (0xfffffffffffd86) *)
- := (Const WO~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~0~0~0~1~1~0).
-Notation "'0xfffffffffffe8a'" (* 72057594037927562 (0xfffffffffffe8a) *)
- := (Const 72057594037927562%Z).
-Notation "'0xfffffffffffe8a'" (* 72057594037927562 (0xfffffffffffe8a) *)
- := (Const WO~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~0~1~0~0~0~1~0~1~0).
-Notation "'0xffffffffffff2e'" (* 72057594037927726 (0xffffffffffff2e) *)
- := (Const 72057594037927726%Z).
-Notation "'0xffffffffffff2e'" (* 72057594037927726 (0xffffffffffff2e) *)
- := (Const WO~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~0~0~1~0~1~1~1~0).
-Notation "'0xffffffffffff8b'" (* 72057594037927819 (0xffffffffffff8b) *)
- := (Const 72057594037927819%Z).
-Notation "'0xffffffffffff8b'" (* 72057594037927819 (0xffffffffffff8b) *)
- := (Const WO~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~0~0~0~1~0~1~1).
-Notation "'0xffffffffffffc2'" (* 72057594037927874 (0xffffffffffffc2) *)
- := (Const 72057594037927874%Z).
-Notation "'0xffffffffffffc2'" (* 72057594037927874 (0xffffffffffffc2) *)
- := (Const WO~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~0~0~0~0~1~0).
-Notation "'0xffffffffffffce'" (* 72057594037927886 (0xffffffffffffce) *)
- := (Const 72057594037927886%Z).
-Notation "'0xffffffffffffce'" (* 72057594037927886 (0xffffffffffffce) *)
- := (Const WO~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~0~0~1~1~1~0).
-Notation "'0xffffffffffffd6'" (* 72057594037927894 (0xffffffffffffd6) *)
- := (Const 72057594037927894%Z).
-Notation "'0xffffffffffffd6'" (* 72057594037927894 (0xffffffffffffd6) *)
- := (Const WO~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~0~1~0~1~1~0).
-Notation "'0xffffffffffffda'" (* 72057594037927898 (0xffffffffffffda) *)
- := (Const 72057594037927898%Z).
-Notation "'0xffffffffffffda'" (* 72057594037927898 (0xffffffffffffda) *)
- := (Const WO~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~0~1~1~0~1~0).
-Notation "'0xffffffffffffde'" (* 72057594037927902 (0xffffffffffffde) *)
- := (Const 72057594037927902%Z).
-Notation "'0xffffffffffffde'" (* 72057594037927902 (0xffffffffffffde) *)
- := (Const WO~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~0~1~1~1~1~0).
-Notation "'0xffffffffffffeb'" (* 72057594037927915 (0xffffffffffffeb) *)
- := (Const 72057594037927915%Z).
-Notation "'0xffffffffffffeb'" (* 72057594037927915 (0xffffffffffffeb) *)
- := (Const WO~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~1~0~1~1).
-Notation "'0xffffffffffffef'" (* 72057594037927919 (0xffffffffffffef) *)
- := (Const 72057594037927919%Z).
-Notation "'0xffffffffffffef'" (* 72057594037927919 (0xffffffffffffef) *)
- := (Const WO~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~1~1~1~1).
-Notation "'0xfffffffffffffb'" (* 72057594037927931 (0xfffffffffffffb) *)
- := (Const 72057594037927931%Z).
-Notation "'0xfffffffffffffb'" (* 72057594037927931 (0xfffffffffffffb) *)
- := (Const WO~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~1~1~0~1~1).
-Notation "'0xfffffffffffffd'" (* 72057594037927933 (0xfffffffffffffd) *)
- := (Const 72057594037927933%Z).
-Notation "'0xfffffffffffffd'" (* 72057594037927933 (0xfffffffffffffd) *)
- := (Const WO~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~1~1~1~0~1).
-Notation "'0xfffffffffffffe'" (* 72057594037927934 (0xfffffffffffffe) *)
- := (Const 72057594037927934%Z).
-Notation "'0xfffffffffffffe'" (* 72057594037927934 (0xfffffffffffffe) *)
- := (Const WO~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~1~1~1~1~0).
-Notation "'0xffffffffffffff'" (* 72057594037927935 (0xffffffffffffff) *)
- := (Const 72057594037927935%Z).
-Notation "'0xffffffffffffff'" (* 72057594037927935 (0xffffffffffffff) *)
- := (Const WO~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~1~1~1~1~1).
-Notation "'0x100000000000000'" (* 72057594037927936 (0x100000000000000) *)
- := (Const 72057594037927936%Z).
-Notation "'0x100000000000000'" (* 72057594037927936 (0x100000000000000) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000001'" (* 72057594037927937 (0x100000000000001) *)
- := (Const 72057594037927937%Z).
-Notation "'0x100000000000001'" (* 72057594037927937 (0x100000000000001) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffff16'" (* 144115188075855638 (0x1ffffffffffff16) *)
- := (Const 144115188075855638%Z).
-Notation "'0x1ffffffffffff16'" (* 144115188075855638 (0x1ffffffffffff16) *)
- := (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~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~0~0~1~0~1~1~0).
-Notation "'0x1ffffffffffffd6'" (* 144115188075855830 (0x1ffffffffffffd6) *)
- := (Const 144115188075855830%Z).
-Notation "'0x1ffffffffffffd6'" (* 144115188075855830 (0x1ffffffffffffd6) *)
- := (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~1~1~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~1~1~0).
-Notation "'0x1ffffffffffffde'" (* 144115188075855838 (0x1ffffffffffffde) *)
- := (Const 144115188075855838%Z).
-Notation "'0x1ffffffffffffde'" (* 144115188075855838 (0x1ffffffffffffde) *)
- := (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~1~1~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~1~1~0).
-Notation "'0x1ffffffffffffed'" (* 144115188075855853 (0x1ffffffffffffed) *)
- := (Const 144115188075855853%Z).
-Notation "'0x1ffffffffffffed'" (* 144115188075855853 (0x1ffffffffffffed) *)
- := (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~1~1~1~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).
-Notation "'0x1fffffffffffff1'" (* 144115188075855857 (0x1fffffffffffff1) *)
- := (Const 144115188075855857%Z).
-Notation "'0x1fffffffffffff1'" (* 144115188075855857 (0x1fffffffffffff1) *)
- := (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~1~1~1~1~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~0~0~1).
-Notation "'0x1fffffffffffff6'" (* 144115188075855862 (0x1fffffffffffff6) *)
- := (Const 144115188075855862%Z).
-Notation "'0x1fffffffffffff6'" (* 144115188075855862 (0x1fffffffffffff6) *)
- := (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~1~1~1~1~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).
-Notation "'0x1fffffffffffff7'" (* 144115188075855863 (0x1fffffffffffff7) *)
- := (Const 144115188075855863%Z).
-Notation "'0x1fffffffffffff7'" (* 144115188075855863 (0x1fffffffffffff7) *)
- := (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~1~1~1~1~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~1).
-Notation "'0x1fffffffffffffa'" (* 144115188075855866 (0x1fffffffffffffa) *)
- := (Const 144115188075855866%Z).
-Notation "'0x1fffffffffffffa'" (* 144115188075855866 (0x1fffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0x1fffffffffffffb'" (* 144115188075855867 (0x1fffffffffffffb) *)
- := (Const 144115188075855867%Z).
-Notation "'0x1fffffffffffffb'" (* 144115188075855867 (0x1fffffffffffffb) *)
- := (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~1~1~1~1~1~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).
-Notation "'0x1fffffffffffffc'" (* 144115188075855868 (0x1fffffffffffffc) *)
- := (Const 144115188075855868%Z).
-Notation "'0x1fffffffffffffc'" (* 144115188075855868 (0x1fffffffffffffc) *)
- := (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~1~1~1~1~1~1~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~0).
-Notation "'0x1fffffffffffffd'" (* 144115188075855869 (0x1fffffffffffffd) *)
- := (Const 144115188075855869%Z).
-Notation "'0x1fffffffffffffd'" (* 144115188075855869 (0x1fffffffffffffd) *)
- := (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~1~1~1~1~1~1~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).
-Notation "'0x1fffffffffffffe'" (* 144115188075855870 (0x1fffffffffffffe) *)
- := (Const 144115188075855870%Z).
-Notation "'0x1fffffffffffffe'" (* 144115188075855870 (0x1fffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0x1ffffffffffffff'" (* 144115188075855871 (0x1ffffffffffffff) *)
- := (Const 144115188075855871%Z).
-Notation "'0x1ffffffffffffff'" (* 144115188075855871 (0x1ffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0x200000000000000'" (* 144115188075855872 (0x200000000000000) *)
- := (Const 144115188075855872%Z).
-Notation "'0x200000000000000'" (* 144115188075855872 (0x200000000000000) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000001'" (* 144115188075855873 (0x200000000000001) *)
- := (Const 144115188075855873%Z).
-Notation "'0x200000000000001'" (* 144115188075855873 (0x200000000000001) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffda'" (* 288230376151711706 (0x3ffffffffffffda) *)
- := (Const 288230376151711706%Z).
-Notation "'0x3ffffffffffffda'" (* 288230376151711706 (0x3ffffffffffffda) *)
- := (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~1~1~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 "'0x3ffffffffffffe1'" (* 288230376151711713 (0x3ffffffffffffe1) *)
- := (Const 288230376151711713%Z).
-Notation "'0x3ffffffffffffe1'" (* 288230376151711713 (0x3ffffffffffffe1) *)
- := (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~1~1~1~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~0~0~0~1).
-Notation "'0x3ffffffffffffe2'" (* 288230376151711714 (0x3ffffffffffffe2) *)
- := (Const 288230376151711714%Z).
-Notation "'0x3ffffffffffffe2'" (* 288230376151711714 (0x3ffffffffffffe2) *)
- := (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~1~1~1~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~0~0~1~0).
-Notation "'0x3ffffffffffffe5'" (* 288230376151711717 (0x3ffffffffffffe5) *)
- := (Const 288230376151711717%Z).
-Notation "'0x3ffffffffffffe5'" (* 288230376151711717 (0x3ffffffffffffe5) *)
- := (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~1~1~1~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~0~1~0~1).
-Notation "'0x3ffffffffffffee'" (* 288230376151711726 (0x3ffffffffffffee) *)
- := (Const 288230376151711726%Z).
-Notation "'0x3ffffffffffffee'" (* 288230376151711726 (0x3ffffffffffffee) *)
- := (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~1~1~1~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~1~0).
-Notation "'0x3ffffffffffffef'" (* 288230376151711727 (0x3ffffffffffffef) *)
- := (Const 288230376151711727%Z).
-Notation "'0x3ffffffffffffef'" (* 288230376151711727 (0x3ffffffffffffef) *)
- := (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~1~1~1~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~1~1).
-Notation "'0x3fffffffffffff6'" (* 288230376151711734 (0x3fffffffffffff6) *)
- := (Const 288230376151711734%Z).
-Notation "'0x3fffffffffffff6'" (* 288230376151711734 (0x3fffffffffffff6) *)
- := (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~1~1~1~1~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).
-Notation "'0x3fffffffffffffa'" (* 288230376151711738 (0x3fffffffffffffa) *)
- := (Const 288230376151711738%Z).
-Notation "'0x3fffffffffffffa'" (* 288230376151711738 (0x3fffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0x3fffffffffffffc'" (* 288230376151711740 (0x3fffffffffffffc) *)
- := (Const 288230376151711740%Z).
-Notation "'0x3fffffffffffffc'" (* 288230376151711740 (0x3fffffffffffffc) *)
- := (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~1~1~1~1~1~1~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~0).
-Notation "'0x3fffffffffffffd'" (* 288230376151711741 (0x3fffffffffffffd) *)
- := (Const 288230376151711741%Z).
-Notation "'0x3fffffffffffffd'" (* 288230376151711741 (0x3fffffffffffffd) *)
- := (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~1~1~1~1~1~1~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).
-Notation "'0x3fffffffffffffe'" (* 288230376151711742 (0x3fffffffffffffe) *)
- := (Const 288230376151711742%Z).
-Notation "'0x3fffffffffffffe'" (* 288230376151711742 (0x3fffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0x3ffffffffffffff'" (* 288230376151711743 (0x3ffffffffffffff) *)
- := (Const 288230376151711743%Z).
-Notation "'0x3ffffffffffffff'" (* 288230376151711743 (0x3ffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0x400000000000000'" (* 288230376151711744 (0x400000000000000) *)
- := (Const 288230376151711744%Z).
-Notation "'0x400000000000000'" (* 288230376151711744 (0x400000000000000) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000001'" (* 288230376151711745 (0x400000000000001) *)
- := (Const 288230376151711745%Z).
-Notation "'0x400000000000001'" (* 288230376151711745 (0x400000000000001) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffc2'" (* 576460752303423426 (0x7ffffffffffffc2) *)
- := (Const 576460752303423426%Z).
-Notation "'0x7ffffffffffffc2'" (* 576460752303423426 (0x7ffffffffffffc2) *)
- := (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~1~1~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~0~0~0~1~0).
-Notation "'0x7ffffffffffffca'" (* 576460752303423434 (0x7ffffffffffffca) *)
- := (Const 576460752303423434%Z).
-Notation "'0x7ffffffffffffca'" (* 576460752303423434 (0x7ffffffffffffca) *)
- := (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~1~1~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~0~1~0~1~0).
-Notation "'0x7ffffffffffffde'" (* 576460752303423454 (0x7ffffffffffffde) *)
- := (Const 576460752303423454%Z).
-Notation "'0x7ffffffffffffde'" (* 576460752303423454 (0x7ffffffffffffde) *)
- := (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~1~1~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~1~1~0).
-Notation "'0x7ffffffffffffeb'" (* 576460752303423467 (0x7ffffffffffffeb) *)
- := (Const 576460752303423467%Z).
-Notation "'0x7ffffffffffffeb'" (* 576460752303423467 (0x7ffffffffffffeb) *)
- := (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~1~1~1~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~1~1).
-Notation "'0x7ffffffffffffed'" (* 576460752303423469 (0x7ffffffffffffed) *)
- := (Const 576460752303423469%Z).
-Notation "'0x7ffffffffffffed'" (* 576460752303423469 (0x7ffffffffffffed) *)
- := (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~1~1~1~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).
-Notation "'0x7ffffffffffffef'" (* 576460752303423471 (0x7ffffffffffffef) *)
- := (Const 576460752303423471%Z).
-Notation "'0x7ffffffffffffef'" (* 576460752303423471 (0x7ffffffffffffef) *)
- := (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~1~1~1~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~1~1).
-Notation "'0x7fffffffffffff1'" (* 576460752303423473 (0x7fffffffffffff1) *)
- := (Const 576460752303423473%Z).
-Notation "'0x7fffffffffffff1'" (* 576460752303423473 (0x7fffffffffffff1) *)
- := (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~1~1~1~1~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~0~0~1).
-Notation "'0x7fffffffffffffa'" (* 576460752303423482 (0x7fffffffffffffa) *)
- := (Const 576460752303423482%Z).
-Notation "'0x7fffffffffffffa'" (* 576460752303423482 (0x7fffffffffffffa) *)
- := (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~1~1~1~1~1~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 "'0x7fffffffffffffe'" (* 576460752303423486 (0x7fffffffffffffe) *)
- := (Const 576460752303423486%Z).
-Notation "'0x7fffffffffffffe'" (* 576460752303423486 (0x7fffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0x7ffffffffffffff'" (* 576460752303423487 (0x7ffffffffffffff) *)
- := (Const 576460752303423487%Z).
-Notation "'0x7ffffffffffffff'" (* 576460752303423487 (0x7ffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0x800000000000000'" (* 576460752303423488 (0x800000000000000) *)
- := (Const 576460752303423488%Z).
-Notation "'0x800000000000000'" (* 576460752303423488 (0x800000000000000) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000001'" (* 576460752303423489 (0x800000000000001) *)
- := (Const 576460752303423489%Z).
-Notation "'0x800000000000001'" (* 576460752303423489 (0x800000000000001) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xf83e0f83e0f83e1'" (* 1117984489315730401 (0xf83e0f83e0f83e1) *)
- := (Const 1117984489315730401%Z).
-Notation "'0xf83e0f83e0f83e1'" (* 1117984489315730401 (0xf83e0f83e0f83e1) *)
- := (Const WO~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~0~1~1~1~1~1~0~0~0~0~1).
-Notation "'0xfffffffffffffd6'" (* 1152921504606846934 (0xfffffffffffffd6) *)
- := (Const 1152921504606846934%Z).
-Notation "'0xfffffffffffffd6'" (* 1152921504606846934 (0xfffffffffffffd6) *)
- := (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~1~1~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~1~1~0).
-Notation "'0xfffffffffffffda'" (* 1152921504606846938 (0xfffffffffffffda) *)
- := (Const 1152921504606846938%Z).
-Notation "'0xfffffffffffffda'" (* 1152921504606846938 (0xfffffffffffffda) *)
- := (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~1~1~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 "'0xfffffffffffffde'" (* 1152921504606846942 (0xfffffffffffffde) *)
- := (Const 1152921504606846942%Z).
-Notation "'0xfffffffffffffde'" (* 1152921504606846942 (0xfffffffffffffde) *)
- := (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~1~1~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~1~1~0).
-Notation "'0xfffffffffffffe2'" (* 1152921504606846946 (0xfffffffffffffe2) *)
- := (Const 1152921504606846946%Z).
-Notation "'0xfffffffffffffe2'" (* 1152921504606846946 (0xfffffffffffffe2) *)
- := (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~1~1~1~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~0~0~1~0).
-Notation "'0xffffffffffffffe'" (* 1152921504606846974 (0xffffffffffffffe) *)
- := (Const 1152921504606846974%Z).
-Notation "'0xffffffffffffffe'" (* 1152921504606846974 (0xffffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~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 "'0xfffffffffffffff'" (* 1152921504606846975 (0xfffffffffffffff) *)
- := (Const 1152921504606846975%Z).
-Notation "'0xfffffffffffffff'" (* 1152921504606846975 (0xfffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0x1000000000000000'" (* 1152921504606846976 (0x1000000000000000) *)
- := (Const 1152921504606846976%Z).
-Notation "'0x1000000000000000'" (* 1152921504606846976 (0x1000000000000000) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000001'" (* 1152921504606846977 (0x1000000000000001) *)
- := (Const 1152921504606846977%Z).
-Notation "'0x1000000000000001'" (* 1152921504606846977 (0x1000000000000001) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffc'" (* 2305843009213693948 (0x1ffffffffffffffc) *)
- := (Const 2305843009213693948%Z).
-Notation "'0x1ffffffffffffffc'" (* 2305843009213693948 (0x1ffffffffffffffc) *)
- := (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~1~1~1~1~1~1~1~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~0).
-Notation "'0x1ffffffffffffffe'" (* 2305843009213693950 (0x1ffffffffffffffe) *)
- := (Const 2305843009213693950%Z).
-Notation "'0x1ffffffffffffffe'" (* 2305843009213693950 (0x1ffffffffffffffe) *)
- := (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~1~1~1~1~1~1~1~1~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 "'0x1fffffffffffffff'" (* 2305843009213693951 (0x1fffffffffffffff) *)
- := (Const 2305843009213693951%Z).
-Notation "'0x1fffffffffffffff'" (* 2305843009213693951 (0x1fffffffffffffff) *)
- := (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~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000'" (* 2305843009213693952 (0x2000000000000000) *)
- := (Const 2305843009213693952%Z).
-Notation "'0x2000000000000000'" (* 2305843009213693952 (0x2000000000000000) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000001'" (* 2305843009213693953 (0x2000000000000001) *)
- := (Const 2305843009213693953%Z).
-Notation "'0x2000000000000001'" (* 2305843009213693953 (0x2000000000000001) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x2e8ba2e8ba2e8ba3'" (* 3353953467947191203 (0x2e8ba2e8ba2e8ba3) *)
- := (Const 3353953467947191203%Z).
-Notation "'0x2e8ba2e8ba2e8ba3'" (* 3353953467947191203 (0x2e8ba2e8ba2e8ba3) *)
- := (Const WO~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~0~1~1~1~0~1~0~0~0~1~1).
-Notation "'0x34f72c234f72c235'" (* 3816567739388183093 (0x34f72c234f72c235) *)
- := (Const 3816567739388183093%Z).
-Notation "'0x34f72c234f72c235'" (* 3816567739388183093 (0x34f72c234f72c235) *)
- := (Const WO~0~0~1~1~0~1~0~0~1~1~1~1~0~1~1~1~0~0~1~0~1~1~0~0~0~0~1~0~0~0~1~1~0~1~0~0~1~1~1~1~0~1~1~1~0~0~1~0~1~1~0~0~0~0~1~0~0~0~1~1~0~1~0~1).
-Notation "'0x3eddffffffffffff'" (* 4530058275181297663 (0x3eddffffffffffff) *)
- := (Const 4530058275181297663%Z).
-Notation "'0x3eddffffffffffff'" (* 4530058275181297663 (0x3eddffffffffffff) *)
- := (Const WO~0~0~1~1~1~1~1~0~1~1~0~1~1~1~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).
-Notation "'0x3f80ffffffffffff'" (* 4575938696385134591 (0x3f80ffffffffffff) *)
- := (Const 4575938696385134591%Z).
-Notation "'0x3f80ffffffffffff'" (* 4575938696385134591 (0x3f80ffffffffffff) *)
- := (Const WO~0~0~1~1~1~1~1~1~1~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).
-Notation "'0x3fffffffffffffff'" (* 4611686018427387903 (0x3fffffffffffffff) *)
- := (Const 4611686018427387903%Z).
-Notation "'0x3fffffffffffffff'" (* 4611686018427387903 (0x3fffffffffffffff) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x4000000000000000'" (* 4611686018427387904 (0x4000000000000000) *)
- := (Const 4611686018427387904%Z).
-Notation "'0x4000000000000000'" (* 4611686018427387904 (0x4000000000000000) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000001'" (* 4611686018427387905 (0x4000000000000001) *)
- := (Const 4611686018427387905%Z).
-Notation "'0x4000000000000001'" (* 4611686018427387905 (0x4000000000000001) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x4ec4ec4ec4ec4ec5'" (* 5675921253449092805 (0x4ec4ec4ec4ec4ec5) *)
- := (Const 5675921253449092805%Z).
-Notation "'0x4ec4ec4ec4ec4ec5'" (* 5675921253449092805 (0x4ec4ec4ec4ec4ec5) *)
- := (Const WO~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~0~1~1~1~0~1~1~0~0~0~1~0~1).
-Notation "'0x7fffffffffffffffL'" (* 9223372036854775807 (0x7fffffffffffffffL) *)
- := (Const 9223372036854775807%Z).
-Notation "'0x7fffffffffffffffL'" (* 9223372036854775807 (0x7fffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x8000000000000000L'" (* 9223372036854775808 (0x8000000000000000L) *)
- := (Const 9223372036854775808%Z).
-Notation "'0x8000000000000000L'" (* 9223372036854775808 (0x8000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000001L'" (* 9223372036854775809 (0x8000000000000001L) *)
- := (Const 9223372036854775809%Z).
-Notation "'0x8000000000000001L'" (* 9223372036854775809 (0x8000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x84bda12f684bda13L'" (* 9564978408590137875 (0x84bda12f684bda13L) *)
- := (Const 9564978408590137875%Z).
-Notation "'0x84bda12f684bda13L'" (* 9564978408590137875 (0x84bda12f684bda13L) *)
- := (Const WO~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~0~1~1~1~1~0~1~1~0~1~0~0~0~0~1~0~0~1~1).
-Notation "'0x86bca1af286bca1bL'" (* 9708812670373448219 (0x86bca1af286bca1bL) *)
- := (Const 9708812670373448219%Z).
-Notation "'0x86bca1af286bca1bL'" (* 9708812670373448219 (0x86bca1af286bca1bL) *)
- := (Const WO~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~0~1~1~1~1~0~0~1~0~1~0~0~0~0~1~1~0~1~1).
-Notation "'0x8a4472fea18a4473L'" (* 9963214713607832691 (0x8a4472fea18a4473L) *)
- := (Const 9963214713607832691%Z).
-Notation "'0x8a4472fea18a4473L'" (* 9963214713607832691 (0x8a4472fea18a4473L) *)
- := (Const WO~1~0~0~0~1~0~1~0~0~1~0~0~0~1~0~0~0~1~1~1~0~0~1~0~1~1~1~1~1~1~1~0~1~0~1~0~0~0~0~1~1~0~0~0~1~0~1~0~0~1~0~0~0~1~0~0~0~1~1~1~0~0~1~1).
-Notation "'0x8e38e38e38e38e39L'" (* 10248191152060862009 (0x8e38e38e38e38e39L) *)
- := (Const 10248191152060862009%Z).
-Notation "'0x8e38e38e38e38e39L'" (* 10248191152060862009 (0x8e38e38e38e38e39L) *)
- := (Const WO~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~0~1~1~1~0~0~1).
-Notation "'0x8f5c28f5c28f5c29L'" (* 10330176681277348905 (0x8f5c28f5c28f5c29L) *)
- := (Const 10330176681277348905%Z).
-Notation "'0x8f5c28f5c28f5c29L'" (* 10330176681277348905 (0x8f5c28f5c28f5c29L) *)
- := (Const WO~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~0~1~1~1~1~0~1~0~1~1~1~0~0~0~0~1~0~1~0~0~1).
-Notation "'0x8fd8fd8fd8fd8fd9L'" (* 10365313336655843289 (0x8fd8fd8fd8fd8fd9L) *)
- := (Const 10365313336655843289%Z).
-Notation "'0x8fd8fd8fd8fd8fd9L'" (* 10365313336655843289 (0x8fd8fd8fd8fd8fd9L) *)
- := (Const WO~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~0~1~1~1~1~1~1~0~1~1~0~0~1).
-Notation "'0xa53fa94fea53fa95L'" (* 11907422100489763477 (0xa53fa94fea53fa95L) *)
- := (Const 11907422100489763477%Z).
-Notation "'0xa53fa94fea53fa95L'" (* 11907422100489763477 (0xa53fa94fea53fa95L) *)
- := (Const WO~1~0~1~0~0~1~0~1~0~0~1~1~1~1~1~1~1~0~1~0~1~0~0~1~0~1~0~0~1~1~1~1~1~1~1~0~1~0~1~0~0~1~0~1~0~0~1~1~1~1~1~1~1~0~1~0~1~0~0~1~0~1~0~1).
-Notation "'0xaa54ffaa54ffaa55L'" (* 12273715991527008853 (0xaa54ffaa54ffaa55L) *)
- := (Const 12273715991527008853%Z).
-Notation "'0xaa54ffaa54ffaa55L'" (* 12273715991527008853 (0xaa54ffaa54ffaa55L) *)
- := (Const WO~1~0~1~0~1~0~1~0~0~1~0~1~0~1~0~0~1~1~1~1~1~1~1~1~1~0~1~0~1~0~1~0~0~1~0~1~0~1~0~0~1~1~1~1~1~1~1~1~1~0~1~0~1~0~1~0~0~1~0~1~0~1~0~1).
-Notation "'0xaaaaaaaaaaaaaaabL'" (* 12297829382473034411 (0xaaaaaaaaaaaaaaabL) *)
- := (Const 12297829382473034411%Z).
-Notation "'0xaaaaaaaaaaaaaaabL'" (* 12297829382473034411 (0xaaaaaaaaaaaaaaabL) *)
- := (Const WO~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~0~1~1).
-Notation "'0xb0ffffffffffffffL'" (* 12754194144713244671 (0xb0ffffffffffffffL) *)
- := (Const 12754194144713244671%Z).
-Notation "'0xb0ffffffffffffffL'" (* 12754194144713244671 (0xb0ffffffffffffffL) *)
- := (Const WO~1~0~1~1~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~1~1~1~1~1).
-Notation "'0xcccccccccccccccdL'" (* 14757395258967641293 (0xcccccccccccccccdL) *)
- := (Const 14757395258967641293%Z).
-Notation "'0xcccccccccccccccdL'" (* 14757395258967641293 (0xcccccccccccccccdL) *)
- := (Const WO~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~0~1~1~0~1).
-Notation "'0xcebeef94fa86fe2dL'" (* 14897608040525528621 (0xcebeef94fa86fe2dL) *)
- := (Const 14897608040525528621%Z).
-Notation "'0xcebeef94fa86fe2dL'" (* 14897608040525528621 (0xcebeef94fa86fe2dL) *)
- := (Const WO~1~1~0~0~1~1~1~0~1~0~1~1~1~1~1~0~1~1~1~0~1~1~1~1~1~0~0~1~0~1~0~0~1~1~1~1~1~0~1~0~1~0~0~0~0~1~1~0~1~1~1~1~1~1~1~0~0~0~1~0~1~1~0~1).
-Notation "'0xcf3cf3cf3cf3cf3dL'" (* 14933078535860113213 (0xcf3cf3cf3cf3cf3dL) *)
- := (Const 14933078535860113213%Z).
-Notation "'0xcf3cf3cf3cf3cf3dL'" (* 14933078535860113213 (0xcf3cf3cf3cf3cf3dL) *)
- := (Const WO~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~0~1~1~1~1~0~1).
-Notation "'0xcfdcfdcfdcfdcfddL'" (* 14978125529935106013 (0xcfdcfdcfdcfdcfddL) *)
- := (Const 14978125529935106013%Z).
-Notation "'0xcfdcfdcfdcfdcfddL'" (* 14978125529935106013 (0xcfdcfdcfdcfdcfddL) *)
- := (Const WO~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~0~1~1~1~1~1~1~0~1~1~1~0~1).
-Notation "'0xd838091dd2253531L'" (* 15580212934572586289 (0xd838091dd2253531L) *)
- := (Const 15580212934572586289%Z).
-Notation "'0xd838091dd2253531L'" (* 15580212934572586289 (0xd838091dd2253531L) *)
- := (Const WO~1~1~0~1~1~0~0~0~0~0~1~1~1~0~0~0~0~0~0~0~1~0~0~1~0~0~0~1~1~1~0~1~1~1~0~1~0~0~1~0~0~0~1~0~0~1~0~1~0~0~1~1~0~1~0~1~0~0~1~1~0~0~0~1).
-Notation "'0xec9e48ae6f71de15L'" (* 17050145153302519317 (0xec9e48ae6f71de15L) *)
- := (Const 17050145153302519317%Z).
-Notation "'0xec9e48ae6f71de15L'" (* 17050145153302519317 (0xec9e48ae6f71de15L) *)
- := (Const WO~1~1~1~0~1~1~0~0~1~0~0~1~1~1~1~0~0~1~0~0~1~0~0~0~1~0~1~0~1~1~1~0~0~1~1~0~1~1~1~1~0~1~1~1~0~0~0~1~1~1~0~1~1~1~1~0~0~0~0~1~0~1~0~1).
-Notation "'0xeeeeeeeeeeeeeeefL'" (* 17216961135462248175 (0xeeeeeeeeeeeeeeefL) *)
- := (Const 17216961135462248175%Z).
-Notation "'0xeeeeeeeeeeeeeeefL'" (* 17216961135462248175 (0xeeeeeeeeeeeeeeefL) *)
- := (Const WO~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~0~1~1~1~1).
-Notation "'0xef7bdef7bdef7bdfL'" (* 17256631552825064415 (0xef7bdef7bdef7bdfL) *)
- := (Const 17256631552825064415%Z).
-Notation "'0xef7bdef7bdef7bdfL'" (* 17256631552825064415 (0xef7bdef7bdef7bdfL) *)
- := (Const WO~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~0~1~1~1~1~1).
-Notation "'0xf0f0f0f0f0f0f0f1L'" (* 17361641481138401521 (0xf0f0f0f0f0f0f0f1L) *)
- := (Const 17361641481138401521%Z).
-Notation "'0xf0f0f0f0f0f0f0f1L'" (* 17361641481138401521 (0xf0f0f0f0f0f0f0f1L) *)
- := (Const WO~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~0~1~1~1~1~0~0~0~1).
-Notation "'0xfe14ffffffffffffL'" (* 18308539860144619519 (0xfe14ffffffffffffL) *)
- := (Const 18308539860144619519%Z).
-Notation "'0xfe14ffffffffffffL'" (* 18308539860144619519 (0xfe14ffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~0~0~0~0~1~0~1~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).
-Notation "'0xffa7ffffffffffffL'" (* 18421974275759013887 (0xffa7ffffffffffffL) *)
- := (Const 18421974275759013887%Z).
-Notation "'0xffa7ffffffffffffL'" (* 18421974275759013887 (0xffa7ffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~0~1~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 "'0xfffaffffffffffffL'" (* 18445336698825998335 (0xfffaffffffffffffL) *)
- := (Const 18445336698825998335%Z).
-Notation "'0xfffaffffffffffffL'" (* 18445336698825998335 (0xfffaffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~0~1~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).
-Notation "'0xfffeffffffffffffL'" (* 18446462598732840959 (0xfffeffffffffffffL) *)
- := (Const 18446462598732840959%Z).
-Notation "'0xfffeffffffffffffL'" (* 18446462598732840959 (0xfffeffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0xffffefffffffffffL'" (* 18446726481523507199 (0xffffefffffffffffL) *)
- := (Const 18446726481523507199%Z).
-Notation "'0xffffefffffffffffL'" (* 18446726481523507199 (0xffffefffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0xfffffffdffffffffL'" (* 18446744065119617023 (0xfffffffdffffffffL) *)
- := (Const 18446744065119617023%Z).
-Notation "'0xfffffffdffffffffL'" (* 18446744065119617023 (0xfffffffdffffffffL) *)
- := (Const WO~1~1~1~1~1~1~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~1~1~1~1~1~1~1~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 "'0xfffffffefffffc2fL'" (* 18446744069414583343 (0xfffffffefffffc2fL) *)
- := (Const 18446744069414583343%Z).
-Notation "'0xfffffffefffffc2fL'" (* 18446744069414583343 (0xfffffffefffffc2fL) *)
- := (Const WO~1~1~1~1~1~1~1~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~0~0~0~0~1~0~1~1~1~1).
-Notation "'0xfffffffeffffffffL'" (* 18446744069414584319 (0xfffffffeffffffffL) *)
- := (Const 18446744069414584319%Z).
-Notation "'0xfffffffeffffffffL'" (* 18446744069414584319 (0xfffffffeffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~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~1~1~1~1~1~1~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 "'0xffffffff00000000L'" (* 18446744069414584320 (0xffffffff00000000L) *)
- := (Const 18446744069414584320%Z).
-Notation "'0xffffffff00000000L'" (* 18446744069414584320 (0xffffffff00000000L) *)
- := (Const WO~1~1~1~1~1~1~1~1~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0xffffffff00000001L'" (* 18446744069414584321 (0xffffffff00000001L) *)
- := (Const 18446744069414584321%Z).
-Notation "'0xffffffff00000001L'" (* 18446744069414584321 (0xffffffff00000001L) *)
- := (Const WO~1~1~1~1~1~1~1~1~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffeffffL'" (* 18446744073709486079 (0xfffffffffffeffffL) *)
- := (Const 18446744073709486079%Z).
-Notation "'0xfffffffffffeffffL'" (* 18446744073709486079 (0xfffffffffffeffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0xfffffffffffffd03L'" (* 18446744073709550851 (0xfffffffffffffd03L) *)
- := (Const 18446744073709550851%Z).
-Notation "'0xfffffffffffffd03L'" (* 18446744073709550851 (0xfffffffffffffd03L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~0~0~1~1).
-Notation "'0xfffffffffffffdc7L'" (* 18446744073709551047 (0xfffffffffffffdc7L) *)
- := (Const 18446744073709551047%Z).
-Notation "'0xfffffffffffffdc7L'" (* 18446744073709551047 (0xfffffffffffffdc7L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~0~0~0~1~1~1).
-Notation "'0xfffffffffffffe1fL'" (* 18446744073709551135 (0xfffffffffffffe1fL) *)
- := (Const 18446744073709551135%Z).
-Notation "'0xfffffffffffffe1fL'" (* 18446744073709551135 (0xfffffffffffffe1fL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1~1~1~1~1).
-Notation "'0xfffffffffffffe5bL'" (* 18446744073709551195 (0xfffffffffffffe5bL) *)
- := (Const 18446744073709551195%Z).
-Notation "'0xfffffffffffffe5bL'" (* 18446744073709551195 (0xfffffffffffffe5bL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~0~1~1~0~1~1).
-Notation "'0xfffffffffffffec3L'" (* 18446744073709551299 (0xfffffffffffffec3L) *)
- := (Const 18446744073709551299%Z).
-Notation "'0xfffffffffffffec3L'" (* 18446744073709551299 (0xfffffffffffffec3L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1~1).
-Notation "'0xffffffffffffff43L'" (* 18446744073709551427 (0xffffffffffffff43L) *)
- := (Const 18446744073709551427%Z).
-Notation "'0xffffffffffffff43L'" (* 18446744073709551427 (0xffffffffffffff43L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1~1).
-Notation "'0xffffffffffffff45L'" (* 18446744073709551429 (0xffffffffffffff45L) *)
- := (Const 18446744073709551429%Z).
-Notation "'0xffffffffffffff45L'" (* 18446744073709551429 (0xffffffffffffff45L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1~0~1).
-Notation "'0xffffffffffffff8bL'" (* 18446744073709551499 (0xffffffffffffff8bL) *)
- := (Const 18446744073709551499%Z).
-Notation "'0xffffffffffffff8bL'" (* 18446744073709551499 (0xffffffffffffff8bL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1~0~1~1).
-Notation "'0xffffffffffffff97L'" (* 18446744073709551511 (0xffffffffffffff97L) *)
- := (Const 18446744073709551511%Z).
-Notation "'0xffffffffffffff97L'" (* 18446744073709551511 (0xffffffffffffff97L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~0~1~1~1).
-Notation "'0xffffffffffffffdfL'" (* 18446744073709551583 (0xffffffffffffffdfL) *)
- := (Const 18446744073709551583%Z).
-Notation "'0xffffffffffffffdfL'" (* 18446744073709551583 (0xffffffffffffffdfL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1~1).
-Notation "'0xffffffffffffffe1L'" (* 18446744073709551585 (0xffffffffffffffe1L) *)
- := (Const 18446744073709551585%Z).
-Notation "'0xffffffffffffffe1L'" (* 18446744073709551585 (0xffffffffffffffe1L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~0~1).
-Notation "'0xffffffffffffffe3L'" (* 18446744073709551587 (0xffffffffffffffe3L) *)
- := (Const 18446744073709551587%Z).
-Notation "'0xffffffffffffffe3L'" (* 18446744073709551587 (0xffffffffffffffe3L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1~1).
-Notation "'0xffffffffffffffe5L'" (* 18446744073709551589 (0xffffffffffffffe5L) *)
- := (Const 18446744073709551589%Z).
-Notation "'0xffffffffffffffe5L'" (* 18446744073709551589 (0xffffffffffffffe5L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~0~1).
-Notation "'0xffffffffffffffe7L'" (* 18446744073709551591 (0xffffffffffffffe7L) *)
- := (Const 18446744073709551591%Z).
-Notation "'0xffffffffffffffe7L'" (* 18446744073709551591 (0xffffffffffffffe7L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~1~1).
-Notation "'0xffffffffffffffebL'" (* 18446744073709551595 (0xffffffffffffffebL) *)
- := (Const 18446744073709551595%Z).
-Notation "'0xffffffffffffffebL'" (* 18446744073709551595 (0xffffffffffffffebL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1).
-Notation "'0xffffffffffffffedL'" (* 18446744073709551597 (0xffffffffffffffedL) *)
- := (Const 18446744073709551597%Z).
-Notation "'0xffffffffffffffedL'" (* 18446744073709551597 (0xffffffffffffffedL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0xffffffffffffffefL'" (* 18446744073709551599 (0xffffffffffffffefL) *)
- := (Const 18446744073709551599%Z).
-Notation "'0xffffffffffffffefL'" (* 18446744073709551599 (0xffffffffffffffefL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1~1).
-Notation "'0xfffffffffffffff1L'" (* 18446744073709551601 (0xfffffffffffffff1L) *)
- := (Const 18446744073709551601%Z).
-Notation "'0xfffffffffffffff1L'" (* 18446744073709551601 (0xfffffffffffffff1L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~0~1).
-Notation "'0xfffffffffffffff3L'" (* 18446744073709551603 (0xfffffffffffffff3L) *)
- := (Const 18446744073709551603%Z).
-Notation "'0xfffffffffffffff3L'" (* 18446744073709551603 (0xfffffffffffffff3L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~0~1~1).
-Notation "'0xfffffffffffffff5L'" (* 18446744073709551605 (0xfffffffffffffff5L) *)
- := (Const 18446744073709551605%Z).
-Notation "'0xfffffffffffffff5L'" (* 18446744073709551605 (0xfffffffffffffff5L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1).
-Notation "'0xfffffffffffffff7L'" (* 18446744073709551607 (0xfffffffffffffff7L) *)
- := (Const 18446744073709551607%Z).
-Notation "'0xfffffffffffffff7L'" (* 18446744073709551607 (0xfffffffffffffff7L) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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~1).
-Notation "'0xfffffffffffffffbL'" (* 18446744073709551611 (0xfffffffffffffffbL) *)
- := (Const 18446744073709551611%Z).
-Notation "'0xfffffffffffffffbL'" (* 18446744073709551611 (0xfffffffffffffffbL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0xfffffffffffffffdL'" (* 18446744073709551613 (0xfffffffffffffffdL) *)
- := (Const 18446744073709551613%Z).
-Notation "'0xfffffffffffffffdL'" (* 18446744073709551613 (0xfffffffffffffffdL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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).
-Notation "'0xfffffffffffffffeL'" (* 18446744073709551614 (0xfffffffffffffffeL) *)
- := (Const 18446744073709551614%Z).
-Notation "'0xfffffffffffffffeL'" (* 18446744073709551614 (0xfffffffffffffffeL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0xffffffffffffffffL'" (* 18446744073709551615 (0xffffffffffffffffL) *)
- := (Const 18446744073709551615%Z).
-Notation "'0xffffffffffffffffL'" (* 18446744073709551615 (0xffffffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000L'" (* 18446744073709551616 (0x10000000000000000L) *)
- := (Const 18446744073709551616%Z).
-Notation "'0x10000000000000000L'" (* 18446744073709551616 (0x10000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000001L'" (* 18446744073709551617 (0x10000000000000001L) *)
- := (Const 18446744073709551617%Z).
-Notation "'0x10000000000000001L'" (* 18446744073709551617 (0x10000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffL'" (* 36893488147419103231 (0x1ffffffffffffffffL) *)
- := (Const 36893488147419103231%Z).
-Notation "'0x1ffffffffffffffffL'" (* 36893488147419103231 (0x1ffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x20000000000000000L'" (* 36893488147419103232 (0x20000000000000000L) *)
- := (Const 36893488147419103232%Z).
-Notation "'0x20000000000000000L'" (* 36893488147419103232 (0x20000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000001L'" (* 36893488147419103233 (0x20000000000000001L) *)
- := (Const 36893488147419103233%Z).
-Notation "'0x20000000000000001L'" (* 36893488147419103233 (0x20000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffL'" (* 73786976294838206463 (0x3ffffffffffffffffL) *)
- := (Const 73786976294838206463%Z).
-Notation "'0x3ffffffffffffffffL'" (* 73786976294838206463 (0x3ffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x40000000000000000L'" (* 73786976294838206464 (0x40000000000000000L) *)
- := (Const 73786976294838206464%Z).
-Notation "'0x40000000000000000L'" (* 73786976294838206464 (0x40000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000001L'" (* 73786976294838206465 (0x40000000000000001L) *)
- := (Const 73786976294838206465%Z).
-Notation "'0x40000000000000001L'" (* 73786976294838206465 (0x40000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffL'" (* 147573952589676412927 (0x7ffffffffffffffffL) *)
- := (Const 147573952589676412927%Z).
-Notation "'0x7ffffffffffffffffL'" (* 147573952589676412927 (0x7ffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x80000000000000000L'" (* 147573952589676412928 (0x80000000000000000L) *)
- := (Const 147573952589676412928%Z).
-Notation "'0x80000000000000000L'" (* 147573952589676412928 (0x80000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000001L'" (* 147573952589676412929 (0x80000000000000001L) *)
- := (Const 147573952589676412929%Z).
-Notation "'0x80000000000000001L'" (* 147573952589676412929 (0x80000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffL'" (* 295147905179352825855 (0xfffffffffffffffffL) *)
- := (Const 295147905179352825855%Z).
-Notation "'0xfffffffffffffffffL'" (* 295147905179352825855 (0xfffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x100000000000000000L'" (* 295147905179352825856 (0x100000000000000000L) *)
- := (Const 295147905179352825856%Z).
-Notation "'0x100000000000000000L'" (* 295147905179352825856 (0x100000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000001L'" (* 295147905179352825857 (0x100000000000000001L) *)
- := (Const 295147905179352825857%Z).
-Notation "'0x100000000000000001L'" (* 295147905179352825857 (0x100000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffL'" (* 590295810358705651711 (0x1fffffffffffffffffL) *)
- := (Const 590295810358705651711%Z).
-Notation "'0x1fffffffffffffffffL'" (* 590295810358705651711 (0x1fffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x200000000000000000L'" (* 590295810358705651712 (0x200000000000000000L) *)
- := (Const 590295810358705651712%Z).
-Notation "'0x200000000000000000L'" (* 590295810358705651712 (0x200000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000001L'" (* 590295810358705651713 (0x200000000000000001L) *)
- := (Const 590295810358705651713%Z).
-Notation "'0x200000000000000001L'" (* 590295810358705651713 (0x200000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffL'" (* 1180591620717411303423 (0x3fffffffffffffffffL) *)
- := (Const 1180591620717411303423%Z).
-Notation "'0x3fffffffffffffffffL'" (* 1180591620717411303423 (0x3fffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x400000000000000000L'" (* 1180591620717411303424 (0x400000000000000000L) *)
- := (Const 1180591620717411303424%Z).
-Notation "'0x400000000000000000L'" (* 1180591620717411303424 (0x400000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000001L'" (* 1180591620717411303425 (0x400000000000000001L) *)
- := (Const 1180591620717411303425%Z).
-Notation "'0x400000000000000001L'" (* 1180591620717411303425 (0x400000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffL'" (* 2361183241434822606847 (0x7fffffffffffffffffL) *)
- := (Const 2361183241434822606847%Z).
-Notation "'0x7fffffffffffffffffL'" (* 2361183241434822606847 (0x7fffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x800000000000000000L'" (* 2361183241434822606848 (0x800000000000000000L) *)
- := (Const 2361183241434822606848%Z).
-Notation "'0x800000000000000000L'" (* 2361183241434822606848 (0x800000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000001L'" (* 2361183241434822606849 (0x800000000000000001L) *)
- := (Const 2361183241434822606849%Z).
-Notation "'0x800000000000000001L'" (* 2361183241434822606849 (0x800000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffL'" (* 4722366482869645213695 (0xffffffffffffffffffL) *)
- := (Const 4722366482869645213695%Z).
-Notation "'0xffffffffffffffffffL'" (* 4722366482869645213695 (0xffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x1000000000000000000L'" (* 4722366482869645213696 (0x1000000000000000000L) *)
- := (Const 4722366482869645213696%Z).
-Notation "'0x1000000000000000000L'" (* 4722366482869645213696 (0x1000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000001L'" (* 4722366482869645213697 (0x1000000000000000001L) *)
- := (Const 4722366482869645213697%Z).
-Notation "'0x1000000000000000001L'" (* 4722366482869645213697 (0x1000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffL'" (* 9444732965739290427391 (0x1ffffffffffffffffffL) *)
- := (Const 9444732965739290427391%Z).
-Notation "'0x1ffffffffffffffffffL'" (* 9444732965739290427391 (0x1ffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x2000000000000000000L'" (* 9444732965739290427392 (0x2000000000000000000L) *)
- := (Const 9444732965739290427392%Z).
-Notation "'0x2000000000000000000L'" (* 9444732965739290427392 (0x2000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000001L'" (* 9444732965739290427393 (0x2000000000000000001L) *)
- := (Const 9444732965739290427393%Z).
-Notation "'0x2000000000000000001L'" (* 9444732965739290427393 (0x2000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffL'" (* 18889465931478580854783 (0x3ffffffffffffffffffL) *)
- := (Const 18889465931478580854783%Z).
-Notation "'0x3ffffffffffffffffffL'" (* 18889465931478580854783 (0x3ffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1).
-Notation "'0x4000000000000000000L'" (* 18889465931478580854784 (0x4000000000000000000L) *)
- := (Const 18889465931478580854784%Z).
-Notation "'0x4000000000000000000L'" (* 18889465931478580854784 (0x4000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000001L'" (* 18889465931478580854785 (0x4000000000000000001L) *)
- := (Const 18889465931478580854785%Z).
-Notation "'0x4000000000000000001L'" (* 18889465931478580854785 (0x4000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffL'" (* 37778931862957161709567 (0x7ffffffffffffffffffL) *)
- := (Const 37778931862957161709567%Z).
-Notation "'0x7ffffffffffffffffffL'" (* 37778931862957161709567 (0x7ffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~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 "'0x8000000000000000000L'" (* 37778931862957161709568 (0x8000000000000000000L) *)
- := (Const 37778931862957161709568%Z).
-Notation "'0x8000000000000000000L'" (* 37778931862957161709568 (0x8000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000001L'" (* 37778931862957161709569 (0x8000000000000000001L) *)
- := (Const 37778931862957161709569%Z).
-Notation "'0x8000000000000000001L'" (* 37778931862957161709569 (0x8000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffL'" (* 75557863725914323419135 (0xfffffffffffffffffffL) *)
- := (Const 75557863725914323419135%Z).
-Notation "'0xfffffffffffffffffffL'" (* 75557863725914323419135 (0xfffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~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 "'0x10000000000000000000L'" (* 75557863725914323419136 (0x10000000000000000000L) *)
- := (Const 75557863725914323419136%Z).
-Notation "'0x10000000000000000000L'" (* 75557863725914323419136 (0x10000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000001L'" (* 75557863725914323419137 (0x10000000000000000001L) *)
- := (Const 75557863725914323419137%Z).
-Notation "'0x10000000000000000001L'" (* 75557863725914323419137 (0x10000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffL'" (* 151115727451828646838271 (0x1fffffffffffffffffffL) *)
- := (Const 151115727451828646838271%Z).
-Notation "'0x1fffffffffffffffffffL'" (* 151115727451828646838271 (0x1fffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~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 "'0x20000000000000000000L'" (* 151115727451828646838272 (0x20000000000000000000L) *)
- := (Const 151115727451828646838272%Z).
-Notation "'0x20000000000000000000L'" (* 151115727451828646838272 (0x20000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000001L'" (* 151115727451828646838273 (0x20000000000000000001L) *)
- := (Const 151115727451828646838273%Z).
-Notation "'0x20000000000000000001L'" (* 151115727451828646838273 (0x20000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffL'" (* 302231454903657293676543 (0x3fffffffffffffffffffL) *)
- := (Const 302231454903657293676543%Z).
-Notation "'0x3fffffffffffffffffffL'" (* 302231454903657293676543 (0x3fffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~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 "'0x40000000000000000000L'" (* 302231454903657293676544 (0x40000000000000000000L) *)
- := (Const 302231454903657293676544%Z).
-Notation "'0x40000000000000000000L'" (* 302231454903657293676544 (0x40000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000001L'" (* 302231454903657293676545 (0x40000000000000000001L) *)
- := (Const 302231454903657293676545%Z).
-Notation "'0x40000000000000000001L'" (* 302231454903657293676545 (0x40000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffL'" (* 604462909807314587353087 (0x7fffffffffffffffffffL) *)
- := (Const 604462909807314587353087%Z).
-Notation "'0x7fffffffffffffffffffL'" (* 604462909807314587353087 (0x7fffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~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 "'0x80000000000000000000L'" (* 604462909807314587353088 (0x80000000000000000000L) *)
- := (Const 604462909807314587353088%Z).
-Notation "'0x80000000000000000000L'" (* 604462909807314587353088 (0x80000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000001L'" (* 604462909807314587353089 (0x80000000000000000001L) *)
- := (Const 604462909807314587353089%Z).
-Notation "'0x80000000000000000001L'" (* 604462909807314587353089 (0x80000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffL'" (* 1208925819614629174706175 (0xffffffffffffffffffffL) *)
- := (Const 1208925819614629174706175%Z).
-Notation "'0xffffffffffffffffffffL'" (* 1208925819614629174706175 (0xffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~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 "'0x100000000000000000000L'" (* 1208925819614629174706176 (0x100000000000000000000L) *)
- := (Const 1208925819614629174706176%Z).
-Notation "'0x100000000000000000000L'" (* 1208925819614629174706176 (0x100000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000001L'" (* 1208925819614629174706177 (0x100000000000000000001L) *)
- := (Const 1208925819614629174706177%Z).
-Notation "'0x100000000000000000001L'" (* 1208925819614629174706177 (0x100000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffL'" (* 2417851639229258349412351 (0x1ffffffffffffffffffffL) *)
- := (Const 2417851639229258349412351%Z).
-Notation "'0x1ffffffffffffffffffffL'" (* 2417851639229258349412351 (0x1ffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~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 "'0x200000000000000000000L'" (* 2417851639229258349412352 (0x200000000000000000000L) *)
- := (Const 2417851639229258349412352%Z).
-Notation "'0x200000000000000000000L'" (* 2417851639229258349412352 (0x200000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000001L'" (* 2417851639229258349412353 (0x200000000000000000001L) *)
- := (Const 2417851639229258349412353%Z).
-Notation "'0x200000000000000000001L'" (* 2417851639229258349412353 (0x200000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffL'" (* 4835703278458516698824703 (0x3ffffffffffffffffffffL) *)
- := (Const 4835703278458516698824703%Z).
-Notation "'0x3ffffffffffffffffffffL'" (* 4835703278458516698824703 (0x3ffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~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 "'0x400000000000000000000L'" (* 4835703278458516698824704 (0x400000000000000000000L) *)
- := (Const 4835703278458516698824704%Z).
-Notation "'0x400000000000000000000L'" (* 4835703278458516698824704 (0x400000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000001L'" (* 4835703278458516698824705 (0x400000000000000000001L) *)
- := (Const 4835703278458516698824705%Z).
-Notation "'0x400000000000000000001L'" (* 4835703278458516698824705 (0x400000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffL'" (* 9671406556917033397649407 (0x7ffffffffffffffffffffL) *)
- := (Const 9671406556917033397649407%Z).
-Notation "'0x7ffffffffffffffffffffL'" (* 9671406556917033397649407 (0x7ffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000L'" (* 9671406556917033397649408 (0x800000000000000000000L) *)
- := (Const 9671406556917033397649408%Z).
-Notation "'0x800000000000000000000L'" (* 9671406556917033397649408 (0x800000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000001L'" (* 9671406556917033397649409 (0x800000000000000000001L) *)
- := (Const 9671406556917033397649409%Z).
-Notation "'0x800000000000000000001L'" (* 9671406556917033397649409 (0x800000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffL'" (* 19342813113834066795298815 (0xfffffffffffffffffffffL) *)
- := (Const 19342813113834066795298815%Z).
-Notation "'0xfffffffffffffffffffffL'" (* 19342813113834066795298815 (0xfffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000L'" (* 19342813113834066795298816 (0x1000000000000000000000L) *)
- := (Const 19342813113834066795298816%Z).
-Notation "'0x1000000000000000000000L'" (* 19342813113834066795298816 (0x1000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000001L'" (* 19342813113834066795298817 (0x1000000000000000000001L) *)
- := (Const 19342813113834066795298817%Z).
-Notation "'0x1000000000000000000001L'" (* 19342813113834066795298817 (0x1000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffL'" (* 38685626227668133590597631 (0x1fffffffffffffffffffffL) *)
- := (Const 38685626227668133590597631%Z).
-Notation "'0x1fffffffffffffffffffffL'" (* 38685626227668133590597631 (0x1fffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000L'" (* 38685626227668133590597632 (0x2000000000000000000000L) *)
- := (Const 38685626227668133590597632%Z).
-Notation "'0x2000000000000000000000L'" (* 38685626227668133590597632 (0x2000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000001L'" (* 38685626227668133590597633 (0x2000000000000000000001L) *)
- := (Const 38685626227668133590597633%Z).
-Notation "'0x2000000000000000000001L'" (* 38685626227668133590597633 (0x2000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffff6L'" (* 77371252455336267181195254 (0x3ffffffffffffffffffff6L) *)
- := (Const 77371252455336267181195254%Z).
-Notation "'0x3ffffffffffffffffffff6L'" (* 77371252455336267181195254 (0x3ffffffffffffffffffff6L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~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).
-Notation "'0x3ffffffffffffffffffffeL'" (* 77371252455336267181195262 (0x3ffffffffffffffffffffeL) *)
- := (Const 77371252455336267181195262%Z).
-Notation "'0x3ffffffffffffffffffffeL'" (* 77371252455336267181195262 (0x3ffffffffffffffffffffeL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~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 "'0x3fffffffffffffffffffffL'" (* 77371252455336267181195263 (0x3fffffffffffffffffffffL) *)
- := (Const 77371252455336267181195263%Z).
-Notation "'0x3fffffffffffffffffffffL'" (* 77371252455336267181195263 (0x3fffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000L'" (* 77371252455336267181195264 (0x4000000000000000000000L) *)
- := (Const 77371252455336267181195264%Z).
-Notation "'0x4000000000000000000000L'" (* 77371252455336267181195264 (0x4000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000001L'" (* 77371252455336267181195265 (0x4000000000000000000001L) *)
- := (Const 77371252455336267181195265%Z).
-Notation "'0x4000000000000000000001L'" (* 77371252455336267181195265 (0x4000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffL'" (* 154742504910672534362390527 (0x7fffffffffffffffffffffL) *)
- := (Const 154742504910672534362390527%Z).
-Notation "'0x7fffffffffffffffffffffL'" (* 154742504910672534362390527 (0x7fffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000L'" (* 154742504910672534362390528 (0x8000000000000000000000L) *)
- := (Const 154742504910672534362390528%Z).
-Notation "'0x8000000000000000000000L'" (* 154742504910672534362390528 (0x8000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000001L'" (* 154742504910672534362390529 (0x8000000000000000000001L) *)
- := (Const 154742504910672534362390529%Z).
-Notation "'0x8000000000000000000001L'" (* 154742504910672534362390529 (0x8000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffL'" (* 309485009821345068724781055 (0xffffffffffffffffffffffL) *)
- := (Const 309485009821345068724781055%Z).
-Notation "'0xffffffffffffffffffffffL'" (* 309485009821345068724781055 (0xffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000L'" (* 309485009821345068724781056 (0x10000000000000000000000L) *)
- := (Const 309485009821345068724781056%Z).
-Notation "'0x10000000000000000000000L'" (* 309485009821345068724781056 (0x10000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000001L'" (* 309485009821345068724781057 (0x10000000000000000000001L) *)
- := (Const 309485009821345068724781057%Z).
-Notation "'0x10000000000000000000001L'" (* 309485009821345068724781057 (0x10000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffL'" (* 618970019642690137449562111 (0x1ffffffffffffffffffffffL) *)
- := (Const 618970019642690137449562111%Z).
-Notation "'0x1ffffffffffffffffffffffL'" (* 618970019642690137449562111 (0x1ffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000L'" (* 618970019642690137449562112 (0x20000000000000000000000L) *)
- := (Const 618970019642690137449562112%Z).
-Notation "'0x20000000000000000000000L'" (* 618970019642690137449562112 (0x20000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000001L'" (* 618970019642690137449562113 (0x20000000000000000000001L) *)
- := (Const 618970019642690137449562113%Z).
-Notation "'0x20000000000000000000001L'" (* 618970019642690137449562113 (0x20000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffL'" (* 1237940039285380274899124223 (0x3ffffffffffffffffffffffL) *)
- := (Const 1237940039285380274899124223%Z).
-Notation "'0x3ffffffffffffffffffffffL'" (* 1237940039285380274899124223 (0x3ffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000L'" (* 1237940039285380274899124224 (0x40000000000000000000000L) *)
- := (Const 1237940039285380274899124224%Z).
-Notation "'0x40000000000000000000000L'" (* 1237940039285380274899124224 (0x40000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000001L'" (* 1237940039285380274899124225 (0x40000000000000000000001L) *)
- := (Const 1237940039285380274899124225%Z).
-Notation "'0x40000000000000000000001L'" (* 1237940039285380274899124225 (0x40000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffL'" (* 2475880078570760549798248447 (0x7ffffffffffffffffffffffL) *)
- := (Const 2475880078570760549798248447%Z).
-Notation "'0x7ffffffffffffffffffffffL'" (* 2475880078570760549798248447 (0x7ffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000L'" (* 2475880078570760549798248448 (0x80000000000000000000000L) *)
- := (Const 2475880078570760549798248448%Z).
-Notation "'0x80000000000000000000000L'" (* 2475880078570760549798248448 (0x80000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000001L'" (* 2475880078570760549798248449 (0x80000000000000000000001L) *)
- := (Const 2475880078570760549798248449%Z).
-Notation "'0x80000000000000000000001L'" (* 2475880078570760549798248449 (0x80000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffL'" (* 4951760157141521099596496895 (0xfffffffffffffffffffffffL) *)
- := (Const 4951760157141521099596496895%Z).
-Notation "'0xfffffffffffffffffffffffL'" (* 4951760157141521099596496895 (0xfffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000L'" (* 4951760157141521099596496896 (0x100000000000000000000000L) *)
- := (Const 4951760157141521099596496896%Z).
-Notation "'0x100000000000000000000000L'" (* 4951760157141521099596496896 (0x100000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000001L'" (* 4951760157141521099596496897 (0x100000000000000000000001L) *)
- := (Const 4951760157141521099596496897%Z).
-Notation "'0x100000000000000000000001L'" (* 4951760157141521099596496897 (0x100000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffL'" (* 9903520314283042199192993791 (0x1fffffffffffffffffffffffL) *)
- := (Const 9903520314283042199192993791%Z).
-Notation "'0x1fffffffffffffffffffffffL'" (* 9903520314283042199192993791 (0x1fffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000L'" (* 9903520314283042199192993792 (0x200000000000000000000000L) *)
- := (Const 9903520314283042199192993792%Z).
-Notation "'0x200000000000000000000000L'" (* 9903520314283042199192993792 (0x200000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000001L'" (* 9903520314283042199192993793 (0x200000000000000000000001L) *)
- := (Const 9903520314283042199192993793%Z).
-Notation "'0x200000000000000000000001L'" (* 9903520314283042199192993793 (0x200000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffL'" (* 19807040628566084398385987583 (0x3fffffffffffffffffffffffL) *)
- := (Const 19807040628566084398385987583%Z).
-Notation "'0x3fffffffffffffffffffffffL'" (* 19807040628566084398385987583 (0x3fffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000L'" (* 19807040628566084398385987584 (0x400000000000000000000000L) *)
- := (Const 19807040628566084398385987584%Z).
-Notation "'0x400000000000000000000000L'" (* 19807040628566084398385987584 (0x400000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000001L'" (* 19807040628566084398385987585 (0x400000000000000000000001L) *)
- := (Const 19807040628566084398385987585%Z).
-Notation "'0x400000000000000000000001L'" (* 19807040628566084398385987585 (0x400000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffL'" (* 39614081257132168796771975167 (0x7fffffffffffffffffffffffL) *)
- := (Const 39614081257132168796771975167%Z).
-Notation "'0x7fffffffffffffffffffffffL'" (* 39614081257132168796771975167 (0x7fffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000L'" (* 39614081257132168796771975168 (0x800000000000000000000000L) *)
- := (Const 39614081257132168796771975168%Z).
-Notation "'0x800000000000000000000000L'" (* 39614081257132168796771975168 (0x800000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000001L'" (* 39614081257132168796771975169 (0x800000000000000000000001L) *)
- := (Const 39614081257132168796771975169%Z).
-Notation "'0x800000000000000000000001L'" (* 39614081257132168796771975169 (0x800000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffL'" (* 79228162514264337593543950335 (0xffffffffffffffffffffffffL) *)
- := (Const 79228162514264337593543950335%Z).
-Notation "'0xffffffffffffffffffffffffL'" (* 79228162514264337593543950335 (0xffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000L'" (* 79228162514264337593543950336 (0x1000000000000000000000000L) *)
- := (Const 79228162514264337593543950336%Z).
-Notation "'0x1000000000000000000000000L'" (* 79228162514264337593543950336 (0x1000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000001L'" (* 79228162514264337593543950337 (0x1000000000000000000000001L) *)
- := (Const 79228162514264337593543950337%Z).
-Notation "'0x1000000000000000000000001L'" (* 79228162514264337593543950337 (0x1000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffL'" (* 158456325028528675187087900671 (0x1ffffffffffffffffffffffffL) *)
- := (Const 158456325028528675187087900671%Z).
-Notation "'0x1ffffffffffffffffffffffffL'" (* 158456325028528675187087900671 (0x1ffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000L'" (* 158456325028528675187087900672 (0x2000000000000000000000000L) *)
- := (Const 158456325028528675187087900672%Z).
-Notation "'0x2000000000000000000000000L'" (* 158456325028528675187087900672 (0x2000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000001L'" (* 158456325028528675187087900673 (0x2000000000000000000000001L) *)
- := (Const 158456325028528675187087900673%Z).
-Notation "'0x2000000000000000000000001L'" (* 158456325028528675187087900673 (0x2000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffL'" (* 316912650057057350374175801343 (0x3ffffffffffffffffffffffffL) *)
- := (Const 316912650057057350374175801343%Z).
-Notation "'0x3ffffffffffffffffffffffffL'" (* 316912650057057350374175801343 (0x3ffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000L'" (* 316912650057057350374175801344 (0x4000000000000000000000000L) *)
- := (Const 316912650057057350374175801344%Z).
-Notation "'0x4000000000000000000000000L'" (* 316912650057057350374175801344 (0x4000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000001L'" (* 316912650057057350374175801345 (0x4000000000000000000000001L) *)
- := (Const 316912650057057350374175801345%Z).
-Notation "'0x4000000000000000000000001L'" (* 316912650057057350374175801345 (0x4000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffL'" (* 633825300114114700748351602687 (0x7ffffffffffffffffffffffffL) *)
- := (Const 633825300114114700748351602687%Z).
-Notation "'0x7ffffffffffffffffffffffffL'" (* 633825300114114700748351602687 (0x7ffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000L'" (* 633825300114114700748351602688 (0x8000000000000000000000000L) *)
- := (Const 633825300114114700748351602688%Z).
-Notation "'0x8000000000000000000000000L'" (* 633825300114114700748351602688 (0x8000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000001L'" (* 633825300114114700748351602689 (0x8000000000000000000000001L) *)
- := (Const 633825300114114700748351602689%Z).
-Notation "'0x8000000000000000000000001L'" (* 633825300114114700748351602689 (0x8000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffL'" (* 1267650600228229401496703205375 (0xfffffffffffffffffffffffffL) *)
- := (Const 1267650600228229401496703205375%Z).
-Notation "'0xfffffffffffffffffffffffffL'" (* 1267650600228229401496703205375 (0xfffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000L'" (* 1267650600228229401496703205376 (0x10000000000000000000000000L) *)
- := (Const 1267650600228229401496703205376%Z).
-Notation "'0x10000000000000000000000000L'" (* 1267650600228229401496703205376 (0x10000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000001L'" (* 1267650600228229401496703205377 (0x10000000000000000000000001L) *)
- := (Const 1267650600228229401496703205377%Z).
-Notation "'0x10000000000000000000000001L'" (* 1267650600228229401496703205377 (0x10000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffL'" (* 2535301200456458802993406410751 (0x1fffffffffffffffffffffffffL) *)
- := (Const 2535301200456458802993406410751%Z).
-Notation "'0x1fffffffffffffffffffffffffL'" (* 2535301200456458802993406410751 (0x1fffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000L'" (* 2535301200456458802993406410752 (0x20000000000000000000000000L) *)
- := (Const 2535301200456458802993406410752%Z).
-Notation "'0x20000000000000000000000000L'" (* 2535301200456458802993406410752 (0x20000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000001L'" (* 2535301200456458802993406410753 (0x20000000000000000000000001L) *)
- := (Const 2535301200456458802993406410753%Z).
-Notation "'0x20000000000000000000000001L'" (* 2535301200456458802993406410753 (0x20000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffL'" (* 5070602400912917605986812821503 (0x3fffffffffffffffffffffffffL) *)
- := (Const 5070602400912917605986812821503%Z).
-Notation "'0x3fffffffffffffffffffffffffL'" (* 5070602400912917605986812821503 (0x3fffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000L'" (* 5070602400912917605986812821504 (0x40000000000000000000000000L) *)
- := (Const 5070602400912917605986812821504%Z).
-Notation "'0x40000000000000000000000000L'" (* 5070602400912917605986812821504 (0x40000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000001L'" (* 5070602400912917605986812821505 (0x40000000000000000000000001L) *)
- := (Const 5070602400912917605986812821505%Z).
-Notation "'0x40000000000000000000000001L'" (* 5070602400912917605986812821505 (0x40000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffL'" (* 10141204801825835211973625643007 (0x7fffffffffffffffffffffffffL) *)
- := (Const 10141204801825835211973625643007%Z).
-Notation "'0x7fffffffffffffffffffffffffL'" (* 10141204801825835211973625643007 (0x7fffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000L'" (* 10141204801825835211973625643008 (0x80000000000000000000000000L) *)
- := (Const 10141204801825835211973625643008%Z).
-Notation "'0x80000000000000000000000000L'" (* 10141204801825835211973625643008 (0x80000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000001L'" (* 10141204801825835211973625643009 (0x80000000000000000000000001L) *)
- := (Const 10141204801825835211973625643009%Z).
-Notation "'0x80000000000000000000000001L'" (* 10141204801825835211973625643009 (0x80000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffL'" (* 20282409603651670423947251286015 (0xffffffffffffffffffffffffffL) *)
- := (Const 20282409603651670423947251286015%Z).
-Notation "'0xffffffffffffffffffffffffffL'" (* 20282409603651670423947251286015 (0xffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000L'" (* 20282409603651670423947251286016 (0x100000000000000000000000000L) *)
- := (Const 20282409603651670423947251286016%Z).
-Notation "'0x100000000000000000000000000L'" (* 20282409603651670423947251286016 (0x100000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000001L'" (* 20282409603651670423947251286017 (0x100000000000000000000000001L) *)
- := (Const 20282409603651670423947251286017%Z).
-Notation "'0x100000000000000000000000001L'" (* 20282409603651670423947251286017 (0x100000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffL'" (* 40564819207303340847894502572031 (0x1ffffffffffffffffffffffffffL) *)
- := (Const 40564819207303340847894502572031%Z).
-Notation "'0x1ffffffffffffffffffffffffffL'" (* 40564819207303340847894502572031 (0x1ffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000L'" (* 40564819207303340847894502572032 (0x200000000000000000000000000L) *)
- := (Const 40564819207303340847894502572032%Z).
-Notation "'0x200000000000000000000000000L'" (* 40564819207303340847894502572032 (0x200000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000001L'" (* 40564819207303340847894502572033 (0x200000000000000000000000001L) *)
- := (Const 40564819207303340847894502572033%Z).
-Notation "'0x200000000000000000000000001L'" (* 40564819207303340847894502572033 (0x200000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffL'" (* 81129638414606681695789005144063 (0x3ffffffffffffffffffffffffffL) *)
- := (Const 81129638414606681695789005144063%Z).
-Notation "'0x3ffffffffffffffffffffffffffL'" (* 81129638414606681695789005144063 (0x3ffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000L'" (* 81129638414606681695789005144064 (0x400000000000000000000000000L) *)
- := (Const 81129638414606681695789005144064%Z).
-Notation "'0x400000000000000000000000000L'" (* 81129638414606681695789005144064 (0x400000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000001L'" (* 81129638414606681695789005144065 (0x400000000000000000000000001L) *)
- := (Const 81129638414606681695789005144065%Z).
-Notation "'0x400000000000000000000000001L'" (* 81129638414606681695789005144065 (0x400000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffL'" (* 162259276829213363391578010288127 (0x7ffffffffffffffffffffffffffL) *)
- := (Const 162259276829213363391578010288127%Z).
-Notation "'0x7ffffffffffffffffffffffffffL'" (* 162259276829213363391578010288127 (0x7ffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000L'" (* 162259276829213363391578010288128 (0x800000000000000000000000000L) *)
- := (Const 162259276829213363391578010288128%Z).
-Notation "'0x800000000000000000000000000L'" (* 162259276829213363391578010288128 (0x800000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000001L'" (* 162259276829213363391578010288129 (0x800000000000000000000000001L) *)
- := (Const 162259276829213363391578010288129%Z).
-Notation "'0x800000000000000000000000001L'" (* 162259276829213363391578010288129 (0x800000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffL'" (* 324518553658426726783156020576255 (0xfffffffffffffffffffffffffffL) *)
- := (Const 324518553658426726783156020576255%Z).
-Notation "'0xfffffffffffffffffffffffffffL'" (* 324518553658426726783156020576255 (0xfffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000L'" (* 324518553658426726783156020576256 (0x1000000000000000000000000000L) *)
- := (Const 324518553658426726783156020576256%Z).
-Notation "'0x1000000000000000000000000000L'" (* 324518553658426726783156020576256 (0x1000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000001L'" (* 324518553658426726783156020576257 (0x1000000000000000000000000001L) *)
- := (Const 324518553658426726783156020576257%Z).
-Notation "'0x1000000000000000000000000001L'" (* 324518553658426726783156020576257 (0x1000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffL'" (* 649037107316853453566312041152511 (0x1fffffffffffffffffffffffffffL) *)
- := (Const 649037107316853453566312041152511%Z).
-Notation "'0x1fffffffffffffffffffffffffffL'" (* 649037107316853453566312041152511 (0x1fffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000L'" (* 649037107316853453566312041152512 (0x2000000000000000000000000000L) *)
- := (Const 649037107316853453566312041152512%Z).
-Notation "'0x2000000000000000000000000000L'" (* 649037107316853453566312041152512 (0x2000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000001L'" (* 649037107316853453566312041152513 (0x2000000000000000000000000001L) *)
- := (Const 649037107316853453566312041152513%Z).
-Notation "'0x2000000000000000000000000001L'" (* 649037107316853453566312041152513 (0x2000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffL'" (* 1298074214633706907132624082305023 (0x3fffffffffffffffffffffffffffL) *)
- := (Const 1298074214633706907132624082305023%Z).
-Notation "'0x3fffffffffffffffffffffffffffL'" (* 1298074214633706907132624082305023 (0x3fffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000L'" (* 1298074214633706907132624082305024 (0x4000000000000000000000000000L) *)
- := (Const 1298074214633706907132624082305024%Z).
-Notation "'0x4000000000000000000000000000L'" (* 1298074214633706907132624082305024 (0x4000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000001L'" (* 1298074214633706907132624082305025 (0x4000000000000000000000000001L) *)
- := (Const 1298074214633706907132624082305025%Z).
-Notation "'0x4000000000000000000000000001L'" (* 1298074214633706907132624082305025 (0x4000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffL'" (* 2596148429267413814265248164610047 (0x7fffffffffffffffffffffffffffL) *)
- := (Const 2596148429267413814265248164610047%Z).
-Notation "'0x7fffffffffffffffffffffffffffL'" (* 2596148429267413814265248164610047 (0x7fffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000L'" (* 2596148429267413814265248164610048 (0x8000000000000000000000000000L) *)
- := (Const 2596148429267413814265248164610048%Z).
-Notation "'0x8000000000000000000000000000L'" (* 2596148429267413814265248164610048 (0x8000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000001L'" (* 2596148429267413814265248164610049 (0x8000000000000000000000000001L) *)
- := (Const 2596148429267413814265248164610049%Z).
-Notation "'0x8000000000000000000000000001L'" (* 2596148429267413814265248164610049 (0x8000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffL'" (* 5192296858534827628530496329220095 (0xffffffffffffffffffffffffffffL) *)
- := (Const 5192296858534827628530496329220095%Z).
-Notation "'0xffffffffffffffffffffffffffffL'" (* 5192296858534827628530496329220095 (0xffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000L'" (* 5192296858534827628530496329220096 (0x10000000000000000000000000000L) *)
- := (Const 5192296858534827628530496329220096%Z).
-Notation "'0x10000000000000000000000000000L'" (* 5192296858534827628530496329220096 (0x10000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000001L'" (* 5192296858534827628530496329220097 (0x10000000000000000000000000001L) *)
- := (Const 5192296858534827628530496329220097%Z).
-Notation "'0x10000000000000000000000000001L'" (* 5192296858534827628530496329220097 (0x10000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffL'" (* 10384593717069655257060992658440191 (0x1ffffffffffffffffffffffffffffL) *)
- := (Const 10384593717069655257060992658440191%Z).
-Notation "'0x1ffffffffffffffffffffffffffffL'" (* 10384593717069655257060992658440191 (0x1ffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000L'" (* 10384593717069655257060992658440192 (0x20000000000000000000000000000L) *)
- := (Const 10384593717069655257060992658440192%Z).
-Notation "'0x20000000000000000000000000000L'" (* 10384593717069655257060992658440192 (0x20000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000001L'" (* 10384593717069655257060992658440193 (0x20000000000000000000000000001L) *)
- := (Const 10384593717069655257060992658440193%Z).
-Notation "'0x20000000000000000000000000001L'" (* 10384593717069655257060992658440193 (0x20000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffL'" (* 20769187434139310514121985316880383 (0x3ffffffffffffffffffffffffffffL) *)
- := (Const 20769187434139310514121985316880383%Z).
-Notation "'0x3ffffffffffffffffffffffffffffL'" (* 20769187434139310514121985316880383 (0x3ffffffffffffffffffffffffffffL) *)
- := (Const WO~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000L'" (* 20769187434139310514121985316880384 (0x40000000000000000000000000000L) *)
- := (Const 20769187434139310514121985316880384%Z).
-Notation "'0x40000000000000000000000000000L'" (* 20769187434139310514121985316880384 (0x40000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000001L'" (* 20769187434139310514121985316880385 (0x40000000000000000000000000001L) *)
- := (Const 20769187434139310514121985316880385%Z).
-Notation "'0x40000000000000000000000000001L'" (* 20769187434139310514121985316880385 (0x40000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffL'" (* 41538374868278621028243970633760767 (0x7ffffffffffffffffffffffffffffL) *)
- := (Const 41538374868278621028243970633760767%Z).
-Notation "'0x7ffffffffffffffffffffffffffffL'" (* 41538374868278621028243970633760767 (0x7ffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000L'" (* 41538374868278621028243970633760768 (0x80000000000000000000000000000L) *)
- := (Const 41538374868278621028243970633760768%Z).
-Notation "'0x80000000000000000000000000000L'" (* 41538374868278621028243970633760768 (0x80000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000001L'" (* 41538374868278621028243970633760769 (0x80000000000000000000000000001L) *)
- := (Const 41538374868278621028243970633760769%Z).
-Notation "'0x80000000000000000000000000001L'" (* 41538374868278621028243970633760769 (0x80000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffL'" (* 83076749736557242056487941267521535 (0xfffffffffffffffffffffffffffffL) *)
- := (Const 83076749736557242056487941267521535%Z).
-Notation "'0xfffffffffffffffffffffffffffffL'" (* 83076749736557242056487941267521535 (0xfffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000L'" (* 83076749736557242056487941267521536 (0x100000000000000000000000000000L) *)
- := (Const 83076749736557242056487941267521536%Z).
-Notation "'0x100000000000000000000000000000L'" (* 83076749736557242056487941267521536 (0x100000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000001L'" (* 83076749736557242056487941267521537 (0x100000000000000000000000000001L) *)
- := (Const 83076749736557242056487941267521537%Z).
-Notation "'0x100000000000000000000000000001L'" (* 83076749736557242056487941267521537 (0x100000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffL'" (* 166153499473114484112975882535043071 (0x1fffffffffffffffffffffffffffffL) *)
- := (Const 166153499473114484112975882535043071%Z).
-Notation "'0x1fffffffffffffffffffffffffffffL'" (* 166153499473114484112975882535043071 (0x1fffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000L'" (* 166153499473114484112975882535043072 (0x200000000000000000000000000000L) *)
- := (Const 166153499473114484112975882535043072%Z).
-Notation "'0x200000000000000000000000000000L'" (* 166153499473114484112975882535043072 (0x200000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000001L'" (* 166153499473114484112975882535043073 (0x200000000000000000000000000001L) *)
- := (Const 166153499473114484112975882535043073%Z).
-Notation "'0x200000000000000000000000000001L'" (* 166153499473114484112975882535043073 (0x200000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffL'" (* 332306998946228968225951765070086143 (0x3fffffffffffffffffffffffffffffL) *)
- := (Const 332306998946228968225951765070086143%Z).
-Notation "'0x3fffffffffffffffffffffffffffffL'" (* 332306998946228968225951765070086143 (0x3fffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000L'" (* 332306998946228968225951765070086144 (0x400000000000000000000000000000L) *)
- := (Const 332306998946228968225951765070086144%Z).
-Notation "'0x400000000000000000000000000000L'" (* 332306998946228968225951765070086144 (0x400000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000001L'" (* 332306998946228968225951765070086145 (0x400000000000000000000000000001L) *)
- := (Const 332306998946228968225951765070086145%Z).
-Notation "'0x400000000000000000000000000001L'" (* 332306998946228968225951765070086145 (0x400000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffL'" (* 664613997892457936451903530140172287 (0x7fffffffffffffffffffffffffffffL) *)
- := (Const 664613997892457936451903530140172287%Z).
-Notation "'0x7fffffffffffffffffffffffffffffL'" (* 664613997892457936451903530140172287 (0x7fffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000L'" (* 664613997892457936451903530140172288 (0x800000000000000000000000000000L) *)
- := (Const 664613997892457936451903530140172288%Z).
-Notation "'0x800000000000000000000000000000L'" (* 664613997892457936451903530140172288 (0x800000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000001L'" (* 664613997892457936451903530140172289 (0x800000000000000000000000000001L) *)
- := (Const 664613997892457936451903530140172289%Z).
-Notation "'0x800000000000000000000000000001L'" (* 664613997892457936451903530140172289 (0x800000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffL'" (* 1329227995784915872903807060280344575 (0xffffffffffffffffffffffffffffffL) *)
- := (Const 1329227995784915872903807060280344575%Z).
-Notation "'0xffffffffffffffffffffffffffffffL'" (* 1329227995784915872903807060280344575 (0xffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000L'" (* 1329227995784915872903807060280344576 (0x1000000000000000000000000000000L) *)
- := (Const 1329227995784915872903807060280344576%Z).
-Notation "'0x1000000000000000000000000000000L'" (* 1329227995784915872903807060280344576 (0x1000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000001L'" (* 1329227995784915872903807060280344577 (0x1000000000000000000000000000001L) *)
- := (Const 1329227995784915872903807060280344577%Z).
-Notation "'0x1000000000000000000000000000001L'" (* 1329227995784915872903807060280344577 (0x1000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffL'" (* 2658455991569831745807614120560689151 (0x1ffffffffffffffffffffffffffffffL) *)
- := (Const 2658455991569831745807614120560689151%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffL'" (* 2658455991569831745807614120560689151 (0x1ffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000L'" (* 2658455991569831745807614120560689152 (0x2000000000000000000000000000000L) *)
- := (Const 2658455991569831745807614120560689152%Z).
-Notation "'0x2000000000000000000000000000000L'" (* 2658455991569831745807614120560689152 (0x2000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000001L'" (* 2658455991569831745807614120560689153 (0x2000000000000000000000000000001L) *)
- := (Const 2658455991569831745807614120560689153%Z).
-Notation "'0x2000000000000000000000000000001L'" (* 2658455991569831745807614120560689153 (0x2000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffL'" (* 5316911983139663491615228241121378303 (0x3ffffffffffffffffffffffffffffffL) *)
- := (Const 5316911983139663491615228241121378303%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffL'" (* 5316911983139663491615228241121378303 (0x3ffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000L'" (* 5316911983139663491615228241121378304 (0x4000000000000000000000000000000L) *)
- := (Const 5316911983139663491615228241121378304%Z).
-Notation "'0x4000000000000000000000000000000L'" (* 5316911983139663491615228241121378304 (0x4000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000001L'" (* 5316911983139663491615228241121378305 (0x4000000000000000000000000000001L) *)
- := (Const 5316911983139663491615228241121378305%Z).
-Notation "'0x4000000000000000000000000000001L'" (* 5316911983139663491615228241121378305 (0x4000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffL'" (* 10633823966279326983230456482242756607 (0x7ffffffffffffffffffffffffffffffL) *)
- := (Const 10633823966279326983230456482242756607%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffL'" (* 10633823966279326983230456482242756607 (0x7ffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000L'" (* 10633823966279326983230456482242756608 (0x8000000000000000000000000000000L) *)
- := (Const 10633823966279326983230456482242756608%Z).
-Notation "'0x8000000000000000000000000000000L'" (* 10633823966279326983230456482242756608 (0x8000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000001L'" (* 10633823966279326983230456482242756609 (0x8000000000000000000000000000001L) *)
- := (Const 10633823966279326983230456482242756609%Z).
-Notation "'0x8000000000000000000000000000001L'" (* 10633823966279326983230456482242756609 (0x8000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffL'" (* 21267647932558653966460912964485513215 (0xfffffffffffffffffffffffffffffffL) *)
- := (Const 21267647932558653966460912964485513215%Z).
-Notation "'0xfffffffffffffffffffffffffffffffL'" (* 21267647932558653966460912964485513215 (0xfffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000L'" (* 21267647932558653966460912964485513216 (0x10000000000000000000000000000000L) *)
- := (Const 21267647932558653966460912964485513216%Z).
-Notation "'0x10000000000000000000000000000000L'" (* 21267647932558653966460912964485513216 (0x10000000000000000000000000000000L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000001L'" (* 21267647932558653966460912964485513217 (0x10000000000000000000000000000001L) *)
- := (Const 21267647932558653966460912964485513217%Z).
-Notation "'0x10000000000000000000000000000001L'" (* 21267647932558653966460912964485513217 (0x10000000000000000000000000000001L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffL'" (* 42535295865117307932921825928971026431 (0x1fffffffffffffffffffffffffffffffL) *)
- := (Const 42535295865117307932921825928971026431%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffL'" (* 42535295865117307932921825928971026431 (0x1fffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000L'" (* 42535295865117307932921825928971026432 (0x20000000000000000000000000000000L) *)
- := (Const 42535295865117307932921825928971026432%Z).
-Notation "'0x20000000000000000000000000000000L'" (* 42535295865117307932921825928971026432 (0x20000000000000000000000000000000L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000001L'" (* 42535295865117307932921825928971026433 (0x20000000000000000000000000000001L) *)
- := (Const 42535295865117307932921825928971026433%Z).
-Notation "'0x20000000000000000000000000000001L'" (* 42535295865117307932921825928971026433 (0x20000000000000000000000000000001L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffL'" (* 85070591730234615865843651857942052863 (0x3fffffffffffffffffffffffffffffffL) *)
- := (Const 85070591730234615865843651857942052863%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffL'" (* 85070591730234615865843651857942052863 (0x3fffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000L'" (* 85070591730234615865843651857942052864 (0x40000000000000000000000000000000L) *)
- := (Const 85070591730234615865843651857942052864%Z).
-Notation "'0x40000000000000000000000000000000L'" (* 85070591730234615865843651857942052864 (0x40000000000000000000000000000000L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000001L'" (* 85070591730234615865843651857942052865 (0x40000000000000000000000000000001L) *)
- := (Const 85070591730234615865843651857942052865%Z).
-Notation "'0x40000000000000000000000000000001L'" (* 85070591730234615865843651857942052865 (0x40000000000000000000000000000001L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffL'" (* 170141183460469231731687303715884105727 (0x7fffffffffffffffffffffffffffffffL) *)
- := (Const 170141183460469231731687303715884105727%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffL'" (* 170141183460469231731687303715884105727 (0x7fffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000L'" (* 170141183460469231731687303715884105728 (0x80000000000000000000000000000000L) *)
- := (Const 170141183460469231731687303715884105728%Z).
-Notation "'0x80000000000000000000000000000000L'" (* 170141183460469231731687303715884105728 (0x80000000000000000000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000001L'" (* 170141183460469231731687303715884105729 (0x80000000000000000000000000000001L) *)
- := (Const 170141183460469231731687303715884105729%Z).
-Notation "'0x80000000000000000000000000000001L'" (* 170141183460469231731687303715884105729 (0x80000000000000000000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffff000000010000000000000000L'" (* 340282366841710300967557013911933812736 (0xffffffff000000010000000000000000L) *)
- := (Const 340282366841710300967557013911933812736%Z).
-Notation "'0xffffffff000000010000000000000000L'" (* 340282366841710300967557013911933812736 (0xffffffff000000010000000000000000L) *)
- := (Const WO~1~1~1~1~1~1~1~1~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~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0xffffffffffffffffffffffffffffffffL'" (* 340282366920938463463374607431768211455 (0xffffffffffffffffffffffffffffffffL) *)
- := (Const 340282366920938463463374607431768211455%Z).
-Notation "'0xffffffffffffffffffffffffffffffffL'" (* 340282366920938463463374607431768211455 (0xffffffffffffffffffffffffffffffffL) *)
- := (Const WO~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000L'" (* 340282366920938463463374607431768211456 (0x100000000000000000000000000000000L) *)
- := (Const 340282366920938463463374607431768211456%Z).
-Notation "'0x100000000000000000000000000000000L'" (* 340282366920938463463374607431768211456 (0x100000000000000000000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000001L'" (* 340282366920938463463374607431768211457 (0x100000000000000000000000000000001L) *)
- := (Const 340282366920938463463374607431768211457%Z).
-Notation "'0x100000000000000000000000000000001L'" (* 340282366920938463463374607431768211457 (0x100000000000000000000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffL'" (* 680564733841876926926749214863536422911 (0x1ffffffffffffffffffffffffffffffffL) *)
- := (Const 680564733841876926926749214863536422911%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffL'" (* 680564733841876926926749214863536422911 (0x1ffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000L'" (* 680564733841876926926749214863536422912 (0x200000000000000000000000000000000L) *)
- := (Const 680564733841876926926749214863536422912%Z).
-Notation "'0x200000000000000000000000000000000L'" (* 680564733841876926926749214863536422912 (0x200000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000001L'" (* 680564733841876926926749214863536422913 (0x200000000000000000000000000000001L) *)
- := (Const 680564733841876926926749214863536422913%Z).
-Notation "'0x200000000000000000000000000000001L'" (* 680564733841876926926749214863536422913 (0x200000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffL'" (* 1361129467683753853853498429727072845823 (0x3ffffffffffffffffffffffffffffffffL) *)
- := (Const 1361129467683753853853498429727072845823%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffL'" (* 1361129467683753853853498429727072845823 (0x3ffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000L'" (* 1361129467683753853853498429727072845824 (0x400000000000000000000000000000000L) *)
- := (Const 1361129467683753853853498429727072845824%Z).
-Notation "'0x400000000000000000000000000000000L'" (* 1361129467683753853853498429727072845824 (0x400000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000001L'" (* 1361129467683753853853498429727072845825 (0x400000000000000000000000000000001L) *)
- := (Const 1361129467683753853853498429727072845825%Z).
-Notation "'0x400000000000000000000000000000001L'" (* 1361129467683753853853498429727072845825 (0x400000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffL'" (* 2722258935367507707706996859454145691647 (0x7ffffffffffffffffffffffffffffffffL) *)
- := (Const 2722258935367507707706996859454145691647%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffL'" (* 2722258935367507707706996859454145691647 (0x7ffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000L'" (* 2722258935367507707706996859454145691648 (0x800000000000000000000000000000000L) *)
- := (Const 2722258935367507707706996859454145691648%Z).
-Notation "'0x800000000000000000000000000000000L'" (* 2722258935367507707706996859454145691648 (0x800000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000001L'" (* 2722258935367507707706996859454145691649 (0x800000000000000000000000000000001L) *)
- := (Const 2722258935367507707706996859454145691649%Z).
-Notation "'0x800000000000000000000000000000001L'" (* 2722258935367507707706996859454145691649 (0x800000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffL'" (* 5444517870735015415413993718908291383295 (0xfffffffffffffffffffffffffffffffffL) *)
- := (Const 5444517870735015415413993718908291383295%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffL'" (* 5444517870735015415413993718908291383295 (0xfffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000L'" (* 5444517870735015415413993718908291383296 (0x1000000000000000000000000000000000L) *)
- := (Const 5444517870735015415413993718908291383296%Z).
-Notation "'0x1000000000000000000000000000000000L'" (* 5444517870735015415413993718908291383296 (0x1000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000001L'" (* 5444517870735015415413993718908291383297 (0x1000000000000000000000000000000001L) *)
- := (Const 5444517870735015415413993718908291383297%Z).
-Notation "'0x1000000000000000000000000000000001L'" (* 5444517870735015415413993718908291383297 (0x1000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffL'" (* 10889035741470030830827987437816582766591 (0x1fffffffffffffffffffffffffffffffffL) *)
- := (Const 10889035741470030830827987437816582766591%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffL'" (* 10889035741470030830827987437816582766591 (0x1fffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000L'" (* 10889035741470030830827987437816582766592 (0x2000000000000000000000000000000000L) *)
- := (Const 10889035741470030830827987437816582766592%Z).
-Notation "'0x2000000000000000000000000000000000L'" (* 10889035741470030830827987437816582766592 (0x2000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000001L'" (* 10889035741470030830827987437816582766593 (0x2000000000000000000000000000000001L) *)
- := (Const 10889035741470030830827987437816582766593%Z).
-Notation "'0x2000000000000000000000000000000001L'" (* 10889035741470030830827987437816582766593 (0x2000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffL'" (* 21778071482940061661655974875633165533183 (0x3fffffffffffffffffffffffffffffffffL) *)
- := (Const 21778071482940061661655974875633165533183%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffL'" (* 21778071482940061661655974875633165533183 (0x3fffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000L'" (* 21778071482940061661655974875633165533184 (0x4000000000000000000000000000000000L) *)
- := (Const 21778071482940061661655974875633165533184%Z).
-Notation "'0x4000000000000000000000000000000000L'" (* 21778071482940061661655974875633165533184 (0x4000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000001L'" (* 21778071482940061661655974875633165533185 (0x4000000000000000000000000000000001L) *)
- := (Const 21778071482940061661655974875633165533185%Z).
-Notation "'0x4000000000000000000000000000000001L'" (* 21778071482940061661655974875633165533185 (0x4000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffL'" (* 43556142965880123323311949751266331066367 (0x7fffffffffffffffffffffffffffffffffL) *)
- := (Const 43556142965880123323311949751266331066367%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffL'" (* 43556142965880123323311949751266331066367 (0x7fffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000L'" (* 43556142965880123323311949751266331066368 (0x8000000000000000000000000000000000L) *)
- := (Const 43556142965880123323311949751266331066368%Z).
-Notation "'0x8000000000000000000000000000000000L'" (* 43556142965880123323311949751266331066368 (0x8000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000001L'" (* 43556142965880123323311949751266331066369 (0x8000000000000000000000000000000001L) *)
- := (Const 43556142965880123323311949751266331066369%Z).
-Notation "'0x8000000000000000000000000000000001L'" (* 43556142965880123323311949751266331066369 (0x8000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffL'" (* 87112285931760246646623899502532662132735 (0xffffffffffffffffffffffffffffffffffL) *)
- := (Const 87112285931760246646623899502532662132735%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffL'" (* 87112285931760246646623899502532662132735 (0xffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000L'" (* 87112285931760246646623899502532662132736 (0x10000000000000000000000000000000000L) *)
- := (Const 87112285931760246646623899502532662132736%Z).
-Notation "'0x10000000000000000000000000000000000L'" (* 87112285931760246646623899502532662132736 (0x10000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000001L'" (* 87112285931760246646623899502532662132737 (0x10000000000000000000000000000000001L) *)
- := (Const 87112285931760246646623899502532662132737%Z).
-Notation "'0x10000000000000000000000000000000001L'" (* 87112285931760246646623899502532662132737 (0x10000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffL'" (* 174224571863520493293247799005065324265471 (0x1ffffffffffffffffffffffffffffffffffL) *)
- := (Const 174224571863520493293247799005065324265471%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffL'" (* 174224571863520493293247799005065324265471 (0x1ffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000L'" (* 174224571863520493293247799005065324265472 (0x20000000000000000000000000000000000L) *)
- := (Const 174224571863520493293247799005065324265472%Z).
-Notation "'0x20000000000000000000000000000000000L'" (* 174224571863520493293247799005065324265472 (0x20000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000001L'" (* 174224571863520493293247799005065324265473 (0x20000000000000000000000000000000001L) *)
- := (Const 174224571863520493293247799005065324265473%Z).
-Notation "'0x20000000000000000000000000000000001L'" (* 174224571863520493293247799005065324265473 (0x20000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffL'" (* 348449143727040986586495598010130648530943 (0x3ffffffffffffffffffffffffffffffffffL) *)
- := (Const 348449143727040986586495598010130648530943%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffL'" (* 348449143727040986586495598010130648530943 (0x3ffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000L'" (* 348449143727040986586495598010130648530944 (0x40000000000000000000000000000000000L) *)
- := (Const 348449143727040986586495598010130648530944%Z).
-Notation "'0x40000000000000000000000000000000000L'" (* 348449143727040986586495598010130648530944 (0x40000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000001L'" (* 348449143727040986586495598010130648530945 (0x40000000000000000000000000000000001L) *)
- := (Const 348449143727040986586495598010130648530945%Z).
-Notation "'0x40000000000000000000000000000000001L'" (* 348449143727040986586495598010130648530945 (0x40000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffL'" (* 696898287454081973172991196020261297061887 (0x7ffffffffffffffffffffffffffffffffffL) *)
- := (Const 696898287454081973172991196020261297061887%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffL'" (* 696898287454081973172991196020261297061887 (0x7ffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000L'" (* 696898287454081973172991196020261297061888 (0x80000000000000000000000000000000000L) *)
- := (Const 696898287454081973172991196020261297061888%Z).
-Notation "'0x80000000000000000000000000000000000L'" (* 696898287454081973172991196020261297061888 (0x80000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000001L'" (* 696898287454081973172991196020261297061889 (0x80000000000000000000000000000000001L) *)
- := (Const 696898287454081973172991196020261297061889%Z).
-Notation "'0x80000000000000000000000000000000001L'" (* 696898287454081973172991196020261297061889 (0x80000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffL'" (* 1393796574908163946345982392040522594123775 (0xfffffffffffffffffffffffffffffffffffL) *)
- := (Const 1393796574908163946345982392040522594123775%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffL'" (* 1393796574908163946345982392040522594123775 (0xfffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000L'" (* 1393796574908163946345982392040522594123776 (0x100000000000000000000000000000000000L) *)
- := (Const 1393796574908163946345982392040522594123776%Z).
-Notation "'0x100000000000000000000000000000000000L'" (* 1393796574908163946345982392040522594123776 (0x100000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000001L'" (* 1393796574908163946345982392040522594123777 (0x100000000000000000000000000000000001L) *)
- := (Const 1393796574908163946345982392040522594123777%Z).
-Notation "'0x100000000000000000000000000000000001L'" (* 1393796574908163946345982392040522594123777 (0x100000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffL'" (* 2787593149816327892691964784081045188247551 (0x1fffffffffffffffffffffffffffffffffffL) *)
- := (Const 2787593149816327892691964784081045188247551%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffL'" (* 2787593149816327892691964784081045188247551 (0x1fffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000L'" (* 2787593149816327892691964784081045188247552 (0x200000000000000000000000000000000000L) *)
- := (Const 2787593149816327892691964784081045188247552%Z).
-Notation "'0x200000000000000000000000000000000000L'" (* 2787593149816327892691964784081045188247552 (0x200000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000001L'" (* 2787593149816327892691964784081045188247553 (0x200000000000000000000000000000000001L) *)
- := (Const 2787593149816327892691964784081045188247553%Z).
-Notation "'0x200000000000000000000000000000000001L'" (* 2787593149816327892691964784081045188247553 (0x200000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffL'" (* 5575186299632655785383929568162090376495103 (0x3fffffffffffffffffffffffffffffffffffL) *)
- := (Const 5575186299632655785383929568162090376495103%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffL'" (* 5575186299632655785383929568162090376495103 (0x3fffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000L'" (* 5575186299632655785383929568162090376495104 (0x400000000000000000000000000000000000L) *)
- := (Const 5575186299632655785383929568162090376495104%Z).
-Notation "'0x400000000000000000000000000000000000L'" (* 5575186299632655785383929568162090376495104 (0x400000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000001L'" (* 5575186299632655785383929568162090376495105 (0x400000000000000000000000000000000001L) *)
- := (Const 5575186299632655785383929568162090376495105%Z).
-Notation "'0x400000000000000000000000000000000001L'" (* 5575186299632655785383929568162090376495105 (0x400000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffL'" (* 11150372599265311570767859136324180752990207 (0x7fffffffffffffffffffffffffffffffffffL) *)
- := (Const 11150372599265311570767859136324180752990207%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffL'" (* 11150372599265311570767859136324180752990207 (0x7fffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000L'" (* 11150372599265311570767859136324180752990208 (0x800000000000000000000000000000000000L) *)
- := (Const 11150372599265311570767859136324180752990208%Z).
-Notation "'0x800000000000000000000000000000000000L'" (* 11150372599265311570767859136324180752990208 (0x800000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000001L'" (* 11150372599265311570767859136324180752990209 (0x800000000000000000000000000000000001L) *)
- := (Const 11150372599265311570767859136324180752990209%Z).
-Notation "'0x800000000000000000000000000000000001L'" (* 11150372599265311570767859136324180752990209 (0x800000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffL'" (* 22300745198530623141535718272648361505980415 (0xffffffffffffffffffffffffffffffffffffL) *)
- := (Const 22300745198530623141535718272648361505980415%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffL'" (* 22300745198530623141535718272648361505980415 (0xffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000L'" (* 22300745198530623141535718272648361505980416 (0x1000000000000000000000000000000000000L) *)
- := (Const 22300745198530623141535718272648361505980416%Z).
-Notation "'0x1000000000000000000000000000000000000L'" (* 22300745198530623141535718272648361505980416 (0x1000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000001L'" (* 22300745198530623141535718272648361505980417 (0x1000000000000000000000000000000000001L) *)
- := (Const 22300745198530623141535718272648361505980417%Z).
-Notation "'0x1000000000000000000000000000000000001L'" (* 22300745198530623141535718272648361505980417 (0x1000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffL'" (* 44601490397061246283071436545296723011960831 (0x1ffffffffffffffffffffffffffffffffffffL) *)
- := (Const 44601490397061246283071436545296723011960831%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffL'" (* 44601490397061246283071436545296723011960831 (0x1ffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000L'" (* 44601490397061246283071436545296723011960832 (0x2000000000000000000000000000000000000L) *)
- := (Const 44601490397061246283071436545296723011960832%Z).
-Notation "'0x2000000000000000000000000000000000000L'" (* 44601490397061246283071436545296723011960832 (0x2000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000001L'" (* 44601490397061246283071436545296723011960833 (0x2000000000000000000000000000000000001L) *)
- := (Const 44601490397061246283071436545296723011960833%Z).
-Notation "'0x2000000000000000000000000000000000001L'" (* 44601490397061246283071436545296723011960833 (0x2000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffL'" (* 89202980794122492566142873090593446023921663 (0x3ffffffffffffffffffffffffffffffffffffL) *)
- := (Const 89202980794122492566142873090593446023921663%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffL'" (* 89202980794122492566142873090593446023921663 (0x3ffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000L'" (* 89202980794122492566142873090593446023921664 (0x4000000000000000000000000000000000000L) *)
- := (Const 89202980794122492566142873090593446023921664%Z).
-Notation "'0x4000000000000000000000000000000000000L'" (* 89202980794122492566142873090593446023921664 (0x4000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000001L'" (* 89202980794122492566142873090593446023921665 (0x4000000000000000000000000000000000001L) *)
- := (Const 89202980794122492566142873090593446023921665%Z).
-Notation "'0x4000000000000000000000000000000000001L'" (* 89202980794122492566142873090593446023921665 (0x4000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffL'" (* 178405961588244985132285746181186892047843327 (0x7ffffffffffffffffffffffffffffffffffffL) *)
- := (Const 178405961588244985132285746181186892047843327%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffL'" (* 178405961588244985132285746181186892047843327 (0x7ffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000L'" (* 178405961588244985132285746181186892047843328 (0x8000000000000000000000000000000000000L) *)
- := (Const 178405961588244985132285746181186892047843328%Z).
-Notation "'0x8000000000000000000000000000000000000L'" (* 178405961588244985132285746181186892047843328 (0x8000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000001L'" (* 178405961588244985132285746181186892047843329 (0x8000000000000000000000000000000000001L) *)
- := (Const 178405961588244985132285746181186892047843329%Z).
-Notation "'0x8000000000000000000000000000000000001L'" (* 178405961588244985132285746181186892047843329 (0x8000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffL'" (* 356811923176489970264571492362373784095686655 (0xfffffffffffffffffffffffffffffffffffffL) *)
- := (Const 356811923176489970264571492362373784095686655%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffL'" (* 356811923176489970264571492362373784095686655 (0xfffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000L'" (* 356811923176489970264571492362373784095686656 (0x10000000000000000000000000000000000000L) *)
- := (Const 356811923176489970264571492362373784095686656%Z).
-Notation "'0x10000000000000000000000000000000000000L'" (* 356811923176489970264571492362373784095686656 (0x10000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000001L'" (* 356811923176489970264571492362373784095686657 (0x10000000000000000000000000000000000001L) *)
- := (Const 356811923176489970264571492362373784095686657%Z).
-Notation "'0x10000000000000000000000000000000000001L'" (* 356811923176489970264571492362373784095686657 (0x10000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffL'" (* 713623846352979940529142984724747568191373311 (0x1fffffffffffffffffffffffffffffffffffffL) *)
- := (Const 713623846352979940529142984724747568191373311%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffL'" (* 713623846352979940529142984724747568191373311 (0x1fffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000L'" (* 713623846352979940529142984724747568191373312 (0x20000000000000000000000000000000000000L) *)
- := (Const 713623846352979940529142984724747568191373312%Z).
-Notation "'0x20000000000000000000000000000000000000L'" (* 713623846352979940529142984724747568191373312 (0x20000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000001L'" (* 713623846352979940529142984724747568191373313 (0x20000000000000000000000000000000000001L) *)
- := (Const 713623846352979940529142984724747568191373313%Z).
-Notation "'0x20000000000000000000000000000000000001L'" (* 713623846352979940529142984724747568191373313 (0x20000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffL'" (* 1427247692705959881058285969449495136382746623 (0x3fffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1427247692705959881058285969449495136382746623%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffL'" (* 1427247692705959881058285969449495136382746623 (0x3fffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000L'" (* 1427247692705959881058285969449495136382746624 (0x40000000000000000000000000000000000000L) *)
- := (Const 1427247692705959881058285969449495136382746624%Z).
-Notation "'0x40000000000000000000000000000000000000L'" (* 1427247692705959881058285969449495136382746624 (0x40000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000001L'" (* 1427247692705959881058285969449495136382746625 (0x40000000000000000000000000000000000001L) *)
- := (Const 1427247692705959881058285969449495136382746625%Z).
-Notation "'0x40000000000000000000000000000000000001L'" (* 1427247692705959881058285969449495136382746625 (0x40000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffL'" (* 2854495385411919762116571938898990272765493247 (0x7fffffffffffffffffffffffffffffffffffffL) *)
- := (Const 2854495385411919762116571938898990272765493247%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffL'" (* 2854495385411919762116571938898990272765493247 (0x7fffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000L'" (* 2854495385411919762116571938898990272765493248 (0x80000000000000000000000000000000000000L) *)
- := (Const 2854495385411919762116571938898990272765493248%Z).
-Notation "'0x80000000000000000000000000000000000000L'" (* 2854495385411919762116571938898990272765493248 (0x80000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000001L'" (* 2854495385411919762116571938898990272765493249 (0x80000000000000000000000000000000000001L) *)
- := (Const 2854495385411919762116571938898990272765493249%Z).
-Notation "'0x80000000000000000000000000000000000001L'" (* 2854495385411919762116571938898990272765493249 (0x80000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffL'" (* 5708990770823839524233143877797980545530986495 (0xffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 5708990770823839524233143877797980545530986495%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffL'" (* 5708990770823839524233143877797980545530986495 (0xffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000L'" (* 5708990770823839524233143877797980545530986496 (0x100000000000000000000000000000000000000L) *)
- := (Const 5708990770823839524233143877797980545530986496%Z).
-Notation "'0x100000000000000000000000000000000000000L'" (* 5708990770823839524233143877797980545530986496 (0x100000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000001L'" (* 5708990770823839524233143877797980545530986497 (0x100000000000000000000000000000000000001L) *)
- := (Const 5708990770823839524233143877797980545530986497%Z).
-Notation "'0x100000000000000000000000000000000000001L'" (* 5708990770823839524233143877797980545530986497 (0x100000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffL'" (* 11417981541647679048466287755595961091061972991 (0x1ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 11417981541647679048466287755595961091061972991%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffL'" (* 11417981541647679048466287755595961091061972991 (0x1ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000L'" (* 11417981541647679048466287755595961091061972992 (0x200000000000000000000000000000000000000L) *)
- := (Const 11417981541647679048466287755595961091061972992%Z).
-Notation "'0x200000000000000000000000000000000000000L'" (* 11417981541647679048466287755595961091061972992 (0x200000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000001L'" (* 11417981541647679048466287755595961091061972993 (0x200000000000000000000000000000000000001L) *)
- := (Const 11417981541647679048466287755595961091061972993%Z).
-Notation "'0x200000000000000000000000000000000000001L'" (* 11417981541647679048466287755595961091061972993 (0x200000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffL'" (* 22835963083295358096932575511191922182123945983 (0x3ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 22835963083295358096932575511191922182123945983%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffL'" (* 22835963083295358096932575511191922182123945983 (0x3ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000L'" (* 22835963083295358096932575511191922182123945984 (0x400000000000000000000000000000000000000L) *)
- := (Const 22835963083295358096932575511191922182123945984%Z).
-Notation "'0x400000000000000000000000000000000000000L'" (* 22835963083295358096932575511191922182123945984 (0x400000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000001L'" (* 22835963083295358096932575511191922182123945985 (0x400000000000000000000000000000000000001L) *)
- := (Const 22835963083295358096932575511191922182123945985%Z).
-Notation "'0x400000000000000000000000000000000000001L'" (* 22835963083295358096932575511191922182123945985 (0x400000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffL'" (* 45671926166590716193865151022383844364247891967 (0x7ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 45671926166590716193865151022383844364247891967%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffL'" (* 45671926166590716193865151022383844364247891967 (0x7ffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000L'" (* 45671926166590716193865151022383844364247891968 (0x800000000000000000000000000000000000000L) *)
- := (Const 45671926166590716193865151022383844364247891968%Z).
-Notation "'0x800000000000000000000000000000000000000L'" (* 45671926166590716193865151022383844364247891968 (0x800000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000001L'" (* 45671926166590716193865151022383844364247891969 (0x800000000000000000000000000000000000001L) *)
- := (Const 45671926166590716193865151022383844364247891969%Z).
-Notation "'0x800000000000000000000000000000000000001L'" (* 45671926166590716193865151022383844364247891969 (0x800000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffL'" (* 91343852333181432387730302044767688728495783935 (0xfffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 91343852333181432387730302044767688728495783935%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffL'" (* 91343852333181432387730302044767688728495783935 (0xfffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000L'" (* 91343852333181432387730302044767688728495783936 (0x1000000000000000000000000000000000000000L) *)
- := (Const 91343852333181432387730302044767688728495783936%Z).
-Notation "'0x1000000000000000000000000000000000000000L'" (* 91343852333181432387730302044767688728495783936 (0x1000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000001L'" (* 91343852333181432387730302044767688728495783937 (0x1000000000000000000000000000000000000001L) *)
- := (Const 91343852333181432387730302044767688728495783937%Z).
-Notation "'0x1000000000000000000000000000000000000001L'" (* 91343852333181432387730302044767688728495783937 (0x1000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffL'" (* 182687704666362864775460604089535377456991567871 (0x1fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 182687704666362864775460604089535377456991567871%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffL'" (* 182687704666362864775460604089535377456991567871 (0x1fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000L'" (* 182687704666362864775460604089535377456991567872 (0x2000000000000000000000000000000000000000L) *)
- := (Const 182687704666362864775460604089535377456991567872%Z).
-Notation "'0x2000000000000000000000000000000000000000L'" (* 182687704666362864775460604089535377456991567872 (0x2000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000001L'" (* 182687704666362864775460604089535377456991567873 (0x2000000000000000000000000000000000000001L) *)
- := (Const 182687704666362864775460604089535377456991567873%Z).
-Notation "'0x2000000000000000000000000000000000000001L'" (* 182687704666362864775460604089535377456991567873 (0x2000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffL'" (* 365375409332725729550921208179070754913983135743 (0x3fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 365375409332725729550921208179070754913983135743%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffL'" (* 365375409332725729550921208179070754913983135743 (0x3fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000L'" (* 365375409332725729550921208179070754913983135744 (0x4000000000000000000000000000000000000000L) *)
- := (Const 365375409332725729550921208179070754913983135744%Z).
-Notation "'0x4000000000000000000000000000000000000000L'" (* 365375409332725729550921208179070754913983135744 (0x4000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000001L'" (* 365375409332725729550921208179070754913983135745 (0x4000000000000000000000000000000000000001L) *)
- := (Const 365375409332725729550921208179070754913983135745%Z).
-Notation "'0x4000000000000000000000000000000000000001L'" (* 365375409332725729550921208179070754913983135745 (0x4000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffL'" (* 730750818665451459101842416358141509827966271487 (0x7fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 730750818665451459101842416358141509827966271487%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffL'" (* 730750818665451459101842416358141509827966271487 (0x7fffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000L'" (* 730750818665451459101842416358141509827966271488 (0x8000000000000000000000000000000000000000L) *)
- := (Const 730750818665451459101842416358141509827966271488%Z).
-Notation "'0x8000000000000000000000000000000000000000L'" (* 730750818665451459101842416358141509827966271488 (0x8000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000001L'" (* 730750818665451459101842416358141509827966271489 (0x8000000000000000000000000000000000000001L) *)
- := (Const 730750818665451459101842416358141509827966271489%Z).
-Notation "'0x8000000000000000000000000000000000000001L'" (* 730750818665451459101842416358141509827966271489 (0x8000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffL'" (* 1461501637330902918203684832716283019655932542975 (0xffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1461501637330902918203684832716283019655932542975%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffL'" (* 1461501637330902918203684832716283019655932542975 (0xffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000L'" (* 1461501637330902918203684832716283019655932542976 (0x10000000000000000000000000000000000000000L) *)
- := (Const 1461501637330902918203684832716283019655932542976%Z).
-Notation "'0x10000000000000000000000000000000000000000L'" (* 1461501637330902918203684832716283019655932542976 (0x10000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000001L'" (* 1461501637330902918203684832716283019655932542977 (0x10000000000000000000000000000000000000001L) *)
- := (Const 1461501637330902918203684832716283019655932542977%Z).
-Notation "'0x10000000000000000000000000000000000000001L'" (* 1461501637330902918203684832716283019655932542977 (0x10000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffL'" (* 2923003274661805836407369665432566039311865085951 (0x1ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 2923003274661805836407369665432566039311865085951%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffL'" (* 2923003274661805836407369665432566039311865085951 (0x1ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000L'" (* 2923003274661805836407369665432566039311865085952 (0x20000000000000000000000000000000000000000L) *)
- := (Const 2923003274661805836407369665432566039311865085952%Z).
-Notation "'0x20000000000000000000000000000000000000000L'" (* 2923003274661805836407369665432566039311865085952 (0x20000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000001L'" (* 2923003274661805836407369665432566039311865085953 (0x20000000000000000000000000000000000000001L) *)
- := (Const 2923003274661805836407369665432566039311865085953%Z).
-Notation "'0x20000000000000000000000000000000000000001L'" (* 2923003274661805836407369665432566039311865085953 (0x20000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffL'" (* 5846006549323611672814739330865132078623730171903 (0x3ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 5846006549323611672814739330865132078623730171903%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffL'" (* 5846006549323611672814739330865132078623730171903 (0x3ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000L'" (* 5846006549323611672814739330865132078623730171904 (0x40000000000000000000000000000000000000000L) *)
- := (Const 5846006549323611672814739330865132078623730171904%Z).
-Notation "'0x40000000000000000000000000000000000000000L'" (* 5846006549323611672814739330865132078623730171904 (0x40000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000001L'" (* 5846006549323611672814739330865132078623730171905 (0x40000000000000000000000000000000000000001L) *)
- := (Const 5846006549323611672814739330865132078623730171905%Z).
-Notation "'0x40000000000000000000000000000000000000001L'" (* 5846006549323611672814739330865132078623730171905 (0x40000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffL'" (* 11692013098647223345629478661730264157247460343807 (0x7ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 11692013098647223345629478661730264157247460343807%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffL'" (* 11692013098647223345629478661730264157247460343807 (0x7ffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000L'" (* 11692013098647223345629478661730264157247460343808 (0x80000000000000000000000000000000000000000L) *)
- := (Const 11692013098647223345629478661730264157247460343808%Z).
-Notation "'0x80000000000000000000000000000000000000000L'" (* 11692013098647223345629478661730264157247460343808 (0x80000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000001L'" (* 11692013098647223345629478661730264157247460343809 (0x80000000000000000000000000000000000000001L) *)
- := (Const 11692013098647223345629478661730264157247460343809%Z).
-Notation "'0x80000000000000000000000000000000000000001L'" (* 11692013098647223345629478661730264157247460343809 (0x80000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffL'" (* 23384026197294446691258957323460528314494920687615 (0xfffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 23384026197294446691258957323460528314494920687615%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffL'" (* 23384026197294446691258957323460528314494920687615 (0xfffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000L'" (* 23384026197294446691258957323460528314494920687616 (0x100000000000000000000000000000000000000000L) *)
- := (Const 23384026197294446691258957323460528314494920687616%Z).
-Notation "'0x100000000000000000000000000000000000000000L'" (* 23384026197294446691258957323460528314494920687616 (0x100000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000001L'" (* 23384026197294446691258957323460528314494920687617 (0x100000000000000000000000000000000000000001L) *)
- := (Const 23384026197294446691258957323460528314494920687617%Z).
-Notation "'0x100000000000000000000000000000000000000001L'" (* 23384026197294446691258957323460528314494920687617 (0x100000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffL'" (* 46768052394588893382517914646921056628989841375231 (0x1fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 46768052394588893382517914646921056628989841375231%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffL'" (* 46768052394588893382517914646921056628989841375231 (0x1fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000L'" (* 46768052394588893382517914646921056628989841375232 (0x200000000000000000000000000000000000000000L) *)
- := (Const 46768052394588893382517914646921056628989841375232%Z).
-Notation "'0x200000000000000000000000000000000000000000L'" (* 46768052394588893382517914646921056628989841375232 (0x200000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000001L'" (* 46768052394588893382517914646921056628989841375233 (0x200000000000000000000000000000000000000001L) *)
- := (Const 46768052394588893382517914646921056628989841375233%Z).
-Notation "'0x200000000000000000000000000000000000000001L'" (* 46768052394588893382517914646921056628989841375233 (0x200000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffL'" (* 93536104789177786765035829293842113257979682750463 (0x3fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 93536104789177786765035829293842113257979682750463%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffL'" (* 93536104789177786765035829293842113257979682750463 (0x3fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000L'" (* 93536104789177786765035829293842113257979682750464 (0x400000000000000000000000000000000000000000L) *)
- := (Const 93536104789177786765035829293842113257979682750464%Z).
-Notation "'0x400000000000000000000000000000000000000000L'" (* 93536104789177786765035829293842113257979682750464 (0x400000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000001L'" (* 93536104789177786765035829293842113257979682750465 (0x400000000000000000000000000000000000000001L) *)
- := (Const 93536104789177786765035829293842113257979682750465%Z).
-Notation "'0x400000000000000000000000000000000000000001L'" (* 93536104789177786765035829293842113257979682750465 (0x400000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffL'" (* 187072209578355573530071658587684226515959365500927 (0x7fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 187072209578355573530071658587684226515959365500927%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffL'" (* 187072209578355573530071658587684226515959365500927 (0x7fffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000L'" (* 187072209578355573530071658587684226515959365500928 (0x800000000000000000000000000000000000000000L) *)
- := (Const 187072209578355573530071658587684226515959365500928%Z).
-Notation "'0x800000000000000000000000000000000000000000L'" (* 187072209578355573530071658587684226515959365500928 (0x800000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000001L'" (* 187072209578355573530071658587684226515959365500929 (0x800000000000000000000000000000000000000001L) *)
- := (Const 187072209578355573530071658587684226515959365500929%Z).
-Notation "'0x800000000000000000000000000000000000000001L'" (* 187072209578355573530071658587684226515959365500929 (0x800000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffL'" (* 374144419156711147060143317175368453031918731001855 (0xffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 374144419156711147060143317175368453031918731001855%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffL'" (* 374144419156711147060143317175368453031918731001855 (0xffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000L'" (* 374144419156711147060143317175368453031918731001856 (0x1000000000000000000000000000000000000000000L) *)
- := (Const 374144419156711147060143317175368453031918731001856%Z).
-Notation "'0x1000000000000000000000000000000000000000000L'" (* 374144419156711147060143317175368453031918731001856 (0x1000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000001L'" (* 374144419156711147060143317175368453031918731001857 (0x1000000000000000000000000000000000000000001L) *)
- := (Const 374144419156711147060143317175368453031918731001857%Z).
-Notation "'0x1000000000000000000000000000000000000000001L'" (* 374144419156711147060143317175368453031918731001857 (0x1000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffL'" (* 748288838313422294120286634350736906063837462003711 (0x1ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 748288838313422294120286634350736906063837462003711%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffL'" (* 748288838313422294120286634350736906063837462003711 (0x1ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000L'" (* 748288838313422294120286634350736906063837462003712 (0x2000000000000000000000000000000000000000000L) *)
- := (Const 748288838313422294120286634350736906063837462003712%Z).
-Notation "'0x2000000000000000000000000000000000000000000L'" (* 748288838313422294120286634350736906063837462003712 (0x2000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000001L'" (* 748288838313422294120286634350736906063837462003713 (0x2000000000000000000000000000000000000000001L) *)
- := (Const 748288838313422294120286634350736906063837462003713%Z).
-Notation "'0x2000000000000000000000000000000000000000001L'" (* 748288838313422294120286634350736906063837462003713 (0x2000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffL'" (* 1496577676626844588240573268701473812127674924007423 (0x3ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1496577676626844588240573268701473812127674924007423%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffL'" (* 1496577676626844588240573268701473812127674924007423 (0x3ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000L'" (* 1496577676626844588240573268701473812127674924007424 (0x4000000000000000000000000000000000000000000L) *)
- := (Const 1496577676626844588240573268701473812127674924007424%Z).
-Notation "'0x4000000000000000000000000000000000000000000L'" (* 1496577676626844588240573268701473812127674924007424 (0x4000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000001L'" (* 1496577676626844588240573268701473812127674924007425 (0x4000000000000000000000000000000000000000001L) *)
- := (Const 1496577676626844588240573268701473812127674924007425%Z).
-Notation "'0x4000000000000000000000000000000000000000001L'" (* 1496577676626844588240573268701473812127674924007425 (0x4000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffL'" (* 2993155353253689176481146537402947624255349848014847 (0x7ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 2993155353253689176481146537402947624255349848014847%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffL'" (* 2993155353253689176481146537402947624255349848014847 (0x7ffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000L'" (* 2993155353253689176481146537402947624255349848014848 (0x8000000000000000000000000000000000000000000L) *)
- := (Const 2993155353253689176481146537402947624255349848014848%Z).
-Notation "'0x8000000000000000000000000000000000000000000L'" (* 2993155353253689176481146537402947624255349848014848 (0x8000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000001L'" (* 2993155353253689176481146537402947624255349848014849 (0x8000000000000000000000000000000000000000001L) *)
- := (Const 2993155353253689176481146537402947624255349848014849%Z).
-Notation "'0x8000000000000000000000000000000000000000001L'" (* 2993155353253689176481146537402947624255349848014849 (0x8000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffL'" (* 5986310706507378352962293074805895248510699696029695 (0xfffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 5986310706507378352962293074805895248510699696029695%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffL'" (* 5986310706507378352962293074805895248510699696029695 (0xfffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000000L'" (* 5986310706507378352962293074805895248510699696029696 (0x10000000000000000000000000000000000000000000L) *)
- := (Const 5986310706507378352962293074805895248510699696029696%Z).
-Notation "'0x10000000000000000000000000000000000000000000L'" (* 5986310706507378352962293074805895248510699696029696 (0x10000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000000001L'" (* 5986310706507378352962293074805895248510699696029697 (0x10000000000000000000000000000000000000000001L) *)
- := (Const 5986310706507378352962293074805895248510699696029697%Z).
-Notation "'0x10000000000000000000000000000000000000000001L'" (* 5986310706507378352962293074805895248510699696029697 (0x10000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffL'" (* 11972621413014756705924586149611790497021399392059391 (0x1fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 11972621413014756705924586149611790497021399392059391%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffL'" (* 11972621413014756705924586149611790497021399392059391 (0x1fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000000L'" (* 11972621413014756705924586149611790497021399392059392 (0x20000000000000000000000000000000000000000000L) *)
- := (Const 11972621413014756705924586149611790497021399392059392%Z).
-Notation "'0x20000000000000000000000000000000000000000000L'" (* 11972621413014756705924586149611790497021399392059392 (0x20000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000000001L'" (* 11972621413014756705924586149611790497021399392059393 (0x20000000000000000000000000000000000000000001L) *)
- := (Const 11972621413014756705924586149611790497021399392059393%Z).
-Notation "'0x20000000000000000000000000000000000000000001L'" (* 11972621413014756705924586149611790497021399392059393 (0x20000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffL'" (* 23945242826029513411849172299223580994042798784118783 (0x3fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 23945242826029513411849172299223580994042798784118783%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffL'" (* 23945242826029513411849172299223580994042798784118783 (0x3fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000000L'" (* 23945242826029513411849172299223580994042798784118784 (0x40000000000000000000000000000000000000000000L) *)
- := (Const 23945242826029513411849172299223580994042798784118784%Z).
-Notation "'0x40000000000000000000000000000000000000000000L'" (* 23945242826029513411849172299223580994042798784118784 (0x40000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000000001L'" (* 23945242826029513411849172299223580994042798784118785 (0x40000000000000000000000000000000000000000001L) *)
- := (Const 23945242826029513411849172299223580994042798784118785%Z).
-Notation "'0x40000000000000000000000000000000000000000001L'" (* 23945242826029513411849172299223580994042798784118785 (0x40000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffL'" (* 47890485652059026823698344598447161988085597568237567 (0x7fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 47890485652059026823698344598447161988085597568237567%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffL'" (* 47890485652059026823698344598447161988085597568237567 (0x7fffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000000L'" (* 47890485652059026823698344598447161988085597568237568 (0x80000000000000000000000000000000000000000000L) *)
- := (Const 47890485652059026823698344598447161988085597568237568%Z).
-Notation "'0x80000000000000000000000000000000000000000000L'" (* 47890485652059026823698344598447161988085597568237568 (0x80000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000000001L'" (* 47890485652059026823698344598447161988085597568237569 (0x80000000000000000000000000000000000000000001L) *)
- := (Const 47890485652059026823698344598447161988085597568237569%Z).
-Notation "'0x80000000000000000000000000000000000000000001L'" (* 47890485652059026823698344598447161988085597568237569 (0x80000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffL'" (* 95780971304118053647396689196894323976171195136475135 (0xffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 95780971304118053647396689196894323976171195136475135%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffL'" (* 95780971304118053647396689196894323976171195136475135 (0xffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000000L'" (* 95780971304118053647396689196894323976171195136475136 (0x100000000000000000000000000000000000000000000L) *)
- := (Const 95780971304118053647396689196894323976171195136475136%Z).
-Notation "'0x100000000000000000000000000000000000000000000L'" (* 95780971304118053647396689196894323976171195136475136 (0x100000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000000001L'" (* 95780971304118053647396689196894323976171195136475137 (0x100000000000000000000000000000000000000000001L) *)
- := (Const 95780971304118053647396689196894323976171195136475137%Z).
-Notation "'0x100000000000000000000000000000000000000000001L'" (* 95780971304118053647396689196894323976171195136475137 (0x100000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffL'" (* 191561942608236107294793378393788647952342390272950271 (0x1ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 191561942608236107294793378393788647952342390272950271%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffL'" (* 191561942608236107294793378393788647952342390272950271 (0x1ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000000L'" (* 191561942608236107294793378393788647952342390272950272 (0x200000000000000000000000000000000000000000000L) *)
- := (Const 191561942608236107294793378393788647952342390272950272%Z).
-Notation "'0x200000000000000000000000000000000000000000000L'" (* 191561942608236107294793378393788647952342390272950272 (0x200000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000000001L'" (* 191561942608236107294793378393788647952342390272950273 (0x200000000000000000000000000000000000000000001L) *)
- := (Const 191561942608236107294793378393788647952342390272950273%Z).
-Notation "'0x200000000000000000000000000000000000000000001L'" (* 191561942608236107294793378393788647952342390272950273 (0x200000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffL'" (* 383123885216472214589586756787577295904684780545900543 (0x3ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 383123885216472214589586756787577295904684780545900543%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffL'" (* 383123885216472214589586756787577295904684780545900543 (0x3ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000000L'" (* 383123885216472214589586756787577295904684780545900544 (0x400000000000000000000000000000000000000000000L) *)
- := (Const 383123885216472214589586756787577295904684780545900544%Z).
-Notation "'0x400000000000000000000000000000000000000000000L'" (* 383123885216472214589586756787577295904684780545900544 (0x400000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000000001L'" (* 383123885216472214589586756787577295904684780545900545 (0x400000000000000000000000000000000000000000001L) *)
- := (Const 383123885216472214589586756787577295904684780545900545%Z).
-Notation "'0x400000000000000000000000000000000000000000001L'" (* 383123885216472214589586756787577295904684780545900545 (0x400000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffL'" (* 766247770432944429179173513575154591809369561091801087 (0x7ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 766247770432944429179173513575154591809369561091801087%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffL'" (* 766247770432944429179173513575154591809369561091801087 (0x7ffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000000L'" (* 766247770432944429179173513575154591809369561091801088 (0x800000000000000000000000000000000000000000000L) *)
- := (Const 766247770432944429179173513575154591809369561091801088%Z).
-Notation "'0x800000000000000000000000000000000000000000000L'" (* 766247770432944429179173513575154591809369561091801088 (0x800000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000000001L'" (* 766247770432944429179173513575154591809369561091801089 (0x800000000000000000000000000000000000000000001L) *)
- := (Const 766247770432944429179173513575154591809369561091801089%Z).
-Notation "'0x800000000000000000000000000000000000000000001L'" (* 766247770432944429179173513575154591809369561091801089 (0x800000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffL'" (* 1532495540865888858358347027150309183618739122183602175 (0xfffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1532495540865888858358347027150309183618739122183602175%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffL'" (* 1532495540865888858358347027150309183618739122183602175 (0xfffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000000L'" (* 1532495540865888858358347027150309183618739122183602176 (0x1000000000000000000000000000000000000000000000L) *)
- := (Const 1532495540865888858358347027150309183618739122183602176%Z).
-Notation "'0x1000000000000000000000000000000000000000000000L'" (* 1532495540865888858358347027150309183618739122183602176 (0x1000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000000001L'" (* 1532495540865888858358347027150309183618739122183602177 (0x1000000000000000000000000000000000000000000001L) *)
- := (Const 1532495540865888858358347027150309183618739122183602177%Z).
-Notation "'0x1000000000000000000000000000000000000000000001L'" (* 1532495540865888858358347027150309183618739122183602177 (0x1000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffL'" (* 3064991081731777716716694054300618367237478244367204351 (0x1fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3064991081731777716716694054300618367237478244367204351%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffL'" (* 3064991081731777716716694054300618367237478244367204351 (0x1fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000000L'" (* 3064991081731777716716694054300618367237478244367204352 (0x2000000000000000000000000000000000000000000000L) *)
- := (Const 3064991081731777716716694054300618367237478244367204352%Z).
-Notation "'0x2000000000000000000000000000000000000000000000L'" (* 3064991081731777716716694054300618367237478244367204352 (0x2000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000000001L'" (* 3064991081731777716716694054300618367237478244367204353 (0x2000000000000000000000000000000000000000000001L) *)
- := (Const 3064991081731777716716694054300618367237478244367204353%Z).
-Notation "'0x2000000000000000000000000000000000000000000001L'" (* 3064991081731777716716694054300618367237478244367204353 (0x2000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffL'" (* 6129982163463555433433388108601236734474956488734408703 (0x3fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 6129982163463555433433388108601236734474956488734408703%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffL'" (* 6129982163463555433433388108601236734474956488734408703 (0x3fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000000L'" (* 6129982163463555433433388108601236734474956488734408704 (0x4000000000000000000000000000000000000000000000L) *)
- := (Const 6129982163463555433433388108601236734474956488734408704%Z).
-Notation "'0x4000000000000000000000000000000000000000000000L'" (* 6129982163463555433433388108601236734474956488734408704 (0x4000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000000001L'" (* 6129982163463555433433388108601236734474956488734408705 (0x4000000000000000000000000000000000000000000001L) *)
- := (Const 6129982163463555433433388108601236734474956488734408705%Z).
-Notation "'0x4000000000000000000000000000000000000000000001L'" (* 6129982163463555433433388108601236734474956488734408705 (0x4000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffL'" (* 12259964326927110866866776217202473468949912977468817407 (0x7fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 12259964326927110866866776217202473468949912977468817407%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffL'" (* 12259964326927110866866776217202473468949912977468817407 (0x7fffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000000L'" (* 12259964326927110866866776217202473468949912977468817408 (0x8000000000000000000000000000000000000000000000L) *)
- := (Const 12259964326927110866866776217202473468949912977468817408%Z).
-Notation "'0x8000000000000000000000000000000000000000000000L'" (* 12259964326927110866866776217202473468949912977468817408 (0x8000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000000001L'" (* 12259964326927110866866776217202473468949912977468817409 (0x8000000000000000000000000000000000000000000001L) *)
- := (Const 12259964326927110866866776217202473468949912977468817409%Z).
-Notation "'0x8000000000000000000000000000000000000000000001L'" (* 12259964326927110866866776217202473468949912977468817409 (0x8000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffL'" (* 24519928653854221733733552434404946937899825954937634815 (0xffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 24519928653854221733733552434404946937899825954937634815%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffL'" (* 24519928653854221733733552434404946937899825954937634815 (0xffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000000000L'" (* 24519928653854221733733552434404946937899825954937634816 (0x10000000000000000000000000000000000000000000000L) *)
- := (Const 24519928653854221733733552434404946937899825954937634816%Z).
-Notation "'0x10000000000000000000000000000000000000000000000L'" (* 24519928653854221733733552434404946937899825954937634816 (0x10000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000000000001L'" (* 24519928653854221733733552434404946937899825954937634817 (0x10000000000000000000000000000000000000000000001L) *)
- := (Const 24519928653854221733733552434404946937899825954937634817%Z).
-Notation "'0x10000000000000000000000000000000000000000000001L'" (* 24519928653854221733733552434404946937899825954937634817 (0x10000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffL'" (* 49039857307708443467467104868809893875799651909875269631 (0x1ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 49039857307708443467467104868809893875799651909875269631%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffL'" (* 49039857307708443467467104868809893875799651909875269631 (0x1ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000000000L'" (* 49039857307708443467467104868809893875799651909875269632 (0x20000000000000000000000000000000000000000000000L) *)
- := (Const 49039857307708443467467104868809893875799651909875269632%Z).
-Notation "'0x20000000000000000000000000000000000000000000000L'" (* 49039857307708443467467104868809893875799651909875269632 (0x20000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000000000001L'" (* 49039857307708443467467104868809893875799651909875269633 (0x20000000000000000000000000000000000000000000001L) *)
- := (Const 49039857307708443467467104868809893875799651909875269633%Z).
-Notation "'0x20000000000000000000000000000000000000000000001L'" (* 49039857307708443467467104868809893875799651909875269633 (0x20000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffL'" (* 98079714615416886934934209737619787751599303819750539263 (0x3ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 98079714615416886934934209737619787751599303819750539263%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffL'" (* 98079714615416886934934209737619787751599303819750539263 (0x3ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000000000L'" (* 98079714615416886934934209737619787751599303819750539264 (0x40000000000000000000000000000000000000000000000L) *)
- := (Const 98079714615416886934934209737619787751599303819750539264%Z).
-Notation "'0x40000000000000000000000000000000000000000000000L'" (* 98079714615416886934934209737619787751599303819750539264 (0x40000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000000000001L'" (* 98079714615416886934934209737619787751599303819750539265 (0x40000000000000000000000000000000000000000000001L) *)
- := (Const 98079714615416886934934209737619787751599303819750539265%Z).
-Notation "'0x40000000000000000000000000000000000000000000001L'" (* 98079714615416886934934209737619787751599303819750539265 (0x40000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffL'" (* 196159429230833773869868419475239575503198607639501078527 (0x7ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 196159429230833773869868419475239575503198607639501078527%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffL'" (* 196159429230833773869868419475239575503198607639501078527 (0x7ffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000000000L'" (* 196159429230833773869868419475239575503198607639501078528 (0x80000000000000000000000000000000000000000000000L) *)
- := (Const 196159429230833773869868419475239575503198607639501078528%Z).
-Notation "'0x80000000000000000000000000000000000000000000000L'" (* 196159429230833773869868419475239575503198607639501078528 (0x80000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000000000001L'" (* 196159429230833773869868419475239575503198607639501078529 (0x80000000000000000000000000000000000000000000001L) *)
- := (Const 196159429230833773869868419475239575503198607639501078529%Z).
-Notation "'0x80000000000000000000000000000000000000000000001L'" (* 196159429230833773869868419475239575503198607639501078529 (0x80000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffL'" (* 392318858461667547739736838950479151006397215279002157055 (0xfffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 392318858461667547739736838950479151006397215279002157055%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffL'" (* 392318858461667547739736838950479151006397215279002157055 (0xfffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000000000L'" (* 392318858461667547739736838950479151006397215279002157056 (0x100000000000000000000000000000000000000000000000L) *)
- := (Const 392318858461667547739736838950479151006397215279002157056%Z).
-Notation "'0x100000000000000000000000000000000000000000000000L'" (* 392318858461667547739736838950479151006397215279002157056 (0x100000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000000000001L'" (* 392318858461667547739736838950479151006397215279002157057 (0x100000000000000000000000000000000000000000000001L) *)
- := (Const 392318858461667547739736838950479151006397215279002157057%Z).
-Notation "'0x100000000000000000000000000000000000000000000001L'" (* 392318858461667547739736838950479151006397215279002157057 (0x100000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffL'" (* 784637716923335095479473677900958302012794430558004314111 (0x1fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 784637716923335095479473677900958302012794430558004314111%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffL'" (* 784637716923335095479473677900958302012794430558004314111 (0x1fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000000000L'" (* 784637716923335095479473677900958302012794430558004314112 (0x200000000000000000000000000000000000000000000000L) *)
- := (Const 784637716923335095479473677900958302012794430558004314112%Z).
-Notation "'0x200000000000000000000000000000000000000000000000L'" (* 784637716923335095479473677900958302012794430558004314112 (0x200000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000000000001L'" (* 784637716923335095479473677900958302012794430558004314113 (0x200000000000000000000000000000000000000000000001L) *)
- := (Const 784637716923335095479473677900958302012794430558004314113%Z).
-Notation "'0x200000000000000000000000000000000000000000000001L'" (* 784637716923335095479473677900958302012794430558004314113 (0x200000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1569275433846670190958947355801916604025588861116008628223 (0x3fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1569275433846670190958947355801916604025588861116008628223%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1569275433846670190958947355801916604025588861116008628223 (0x3fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000000000L'" (* 1569275433846670190958947355801916604025588861116008628224 (0x400000000000000000000000000000000000000000000000L) *)
- := (Const 1569275433846670190958947355801916604025588861116008628224%Z).
-Notation "'0x400000000000000000000000000000000000000000000000L'" (* 1569275433846670190958947355801916604025588861116008628224 (0x400000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000000000001L'" (* 1569275433846670190958947355801916604025588861116008628225 (0x400000000000000000000000000000000000000000000001L) *)
- := (Const 1569275433846670190958947355801916604025588861116008628225%Z).
-Notation "'0x400000000000000000000000000000000000000000000001L'" (* 1569275433846670190958947355801916604025588861116008628225 (0x400000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3138550867693340381917894711603833208051177722232017256447 (0x7fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3138550867693340381917894711603833208051177722232017256447%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3138550867693340381917894711603833208051177722232017256447 (0x7fffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000000000L'" (* 3138550867693340381917894711603833208051177722232017256448 (0x800000000000000000000000000000000000000000000000L) *)
- := (Const 3138550867693340381917894711603833208051177722232017256448%Z).
-Notation "'0x800000000000000000000000000000000000000000000000L'" (* 3138550867693340381917894711603833208051177722232017256448 (0x800000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000000000001L'" (* 3138550867693340381917894711603833208051177722232017256449 (0x800000000000000000000000000000000000000000000001L) *)
- := (Const 3138550867693340381917894711603833208051177722232017256449%Z).
-Notation "'0x800000000000000000000000000000000000000000000001L'" (* 3138550867693340381917894711603833208051177722232017256449 (0x800000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6277101735386680763835789423207666416102355444464034512895 (0xffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 6277101735386680763835789423207666416102355444464034512895%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6277101735386680763835789423207666416102355444464034512895 (0xffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000000000L'" (* 6277101735386680763835789423207666416102355444464034512896 (0x1000000000000000000000000000000000000000000000000L) *)
- := (Const 6277101735386680763835789423207666416102355444464034512896%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000L'" (* 6277101735386680763835789423207666416102355444464034512896 (0x1000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000000000001L'" (* 6277101735386680763835789423207666416102355444464034512897 (0x1000000000000000000000000000000000000000000000001L) *)
- := (Const 6277101735386680763835789423207666416102355444464034512897%Z).
-Notation "'0x1000000000000000000000000000000000000000000000001L'" (* 6277101735386680763835789423207666416102355444464034512897 (0x1000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 12554203470773361527671578846415332832204710888928069025791 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 12554203470773361527671578846415332832204710888928069025791%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 12554203470773361527671578846415332832204710888928069025791 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000000000L'" (* 12554203470773361527671578846415332832204710888928069025792 (0x2000000000000000000000000000000000000000000000000L) *)
- := (Const 12554203470773361527671578846415332832204710888928069025792%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000L'" (* 12554203470773361527671578846415332832204710888928069025792 (0x2000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000000000001L'" (* 12554203470773361527671578846415332832204710888928069025793 (0x2000000000000000000000000000000000000000000000001L) *)
- := (Const 12554203470773361527671578846415332832204710888928069025793%Z).
-Notation "'0x2000000000000000000000000000000000000000000000001L'" (* 12554203470773361527671578846415332832204710888928069025793 (0x2000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 25108406941546723055343157692830665664409421777856138051583 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 25108406941546723055343157692830665664409421777856138051583%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 25108406941546723055343157692830665664409421777856138051583 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000000000L'" (* 25108406941546723055343157692830665664409421777856138051584 (0x4000000000000000000000000000000000000000000000000L) *)
- := (Const 25108406941546723055343157692830665664409421777856138051584%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000L'" (* 25108406941546723055343157692830665664409421777856138051584 (0x4000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000000000001L'" (* 25108406941546723055343157692830665664409421777856138051585 (0x4000000000000000000000000000000000000000000000001L) *)
- := (Const 25108406941546723055343157692830665664409421777856138051585%Z).
-Notation "'0x4000000000000000000000000000000000000000000000001L'" (* 25108406941546723055343157692830665664409421777856138051585 (0x4000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 50216813883093446110686315385661331328818843555712276103167 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 50216813883093446110686315385661331328818843555712276103167%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 50216813883093446110686315385661331328818843555712276103167 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000000000L'" (* 50216813883093446110686315385661331328818843555712276103168 (0x8000000000000000000000000000000000000000000000000L) *)
- := (Const 50216813883093446110686315385661331328818843555712276103168%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000L'" (* 50216813883093446110686315385661331328818843555712276103168 (0x8000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000000000001L'" (* 50216813883093446110686315385661331328818843555712276103169 (0x8000000000000000000000000000000000000000000000001L) *)
- := (Const 50216813883093446110686315385661331328818843555712276103169%Z).
-Notation "'0x8000000000000000000000000000000000000000000000001L'" (* 50216813883093446110686315385661331328818843555712276103169 (0x8000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 100433627766186892221372630771322662657637687111424552206335 (0xfffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 100433627766186892221372630771322662657637687111424552206335%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 100433627766186892221372630771322662657637687111424552206335 (0xfffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000000000000L'" (* 100433627766186892221372630771322662657637687111424552206336 (0x10000000000000000000000000000000000000000000000000L) *)
- := (Const 100433627766186892221372630771322662657637687111424552206336%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000L'" (* 100433627766186892221372630771322662657637687111424552206336 (0x10000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000000000000001L'" (* 100433627766186892221372630771322662657637687111424552206337 (0x10000000000000000000000000000000000000000000000001L) *)
- := (Const 100433627766186892221372630771322662657637687111424552206337%Z).
-Notation "'0x10000000000000000000000000000000000000000000000001L'" (* 100433627766186892221372630771322662657637687111424552206337 (0x10000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 200867255532373784442745261542645325315275374222849104412671 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 200867255532373784442745261542645325315275374222849104412671%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 200867255532373784442745261542645325315275374222849104412671 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000000000000L'" (* 200867255532373784442745261542645325315275374222849104412672 (0x20000000000000000000000000000000000000000000000000L) *)
- := (Const 200867255532373784442745261542645325315275374222849104412672%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000L'" (* 200867255532373784442745261542645325315275374222849104412672 (0x20000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000000000000001L'" (* 200867255532373784442745261542645325315275374222849104412673 (0x20000000000000000000000000000000000000000000000001L) *)
- := (Const 200867255532373784442745261542645325315275374222849104412673%Z).
-Notation "'0x20000000000000000000000000000000000000000000000001L'" (* 200867255532373784442745261542645325315275374222849104412673 (0x20000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 401734511064747568885490523085290650630550748445698208825343 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 401734511064747568885490523085290650630550748445698208825343%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 401734511064747568885490523085290650630550748445698208825343 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000000000000L'" (* 401734511064747568885490523085290650630550748445698208825344 (0x40000000000000000000000000000000000000000000000000L) *)
- := (Const 401734511064747568885490523085290650630550748445698208825344%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000L'" (* 401734511064747568885490523085290650630550748445698208825344 (0x40000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000000000000001L'" (* 401734511064747568885490523085290650630550748445698208825345 (0x40000000000000000000000000000000000000000000000001L) *)
- := (Const 401734511064747568885490523085290650630550748445698208825345%Z).
-Notation "'0x40000000000000000000000000000000000000000000000001L'" (* 401734511064747568885490523085290650630550748445698208825345 (0x40000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 803469022129495137770981046170581301261101496891396417650687 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 803469022129495137770981046170581301261101496891396417650687%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 803469022129495137770981046170581301261101496891396417650687 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000000000000L'" (* 803469022129495137770981046170581301261101496891396417650688 (0x80000000000000000000000000000000000000000000000000L) *)
- := (Const 803469022129495137770981046170581301261101496891396417650688%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000L'" (* 803469022129495137770981046170581301261101496891396417650688 (0x80000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000000000000001L'" (* 803469022129495137770981046170581301261101496891396417650689 (0x80000000000000000000000000000000000000000000000001L) *)
- := (Const 803469022129495137770981046170581301261101496891396417650689%Z).
-Notation "'0x80000000000000000000000000000000000000000000000001L'" (* 803469022129495137770981046170581301261101496891396417650689 (0x80000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1606938044258990275541962092341162602522202993782792835301375 (0xffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1606938044258990275541962092341162602522202993782792835301375%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1606938044258990275541962092341162602522202993782792835301375 (0xffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000000000000L'" (* 1606938044258990275541962092341162602522202993782792835301376 (0x100000000000000000000000000000000000000000000000000L) *)
- := (Const 1606938044258990275541962092341162602522202993782792835301376%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000L'" (* 1606938044258990275541962092341162602522202993782792835301376 (0x100000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000000000000001L'" (* 1606938044258990275541962092341162602522202993782792835301377 (0x100000000000000000000000000000000000000000000000001L) *)
- := (Const 1606938044258990275541962092341162602522202993782792835301377%Z).
-Notation "'0x100000000000000000000000000000000000000000000000001L'" (* 1606938044258990275541962092341162602522202993782792835301377 (0x100000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3213876088517980551083924184682325205044405987565585670602751 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3213876088517980551083924184682325205044405987565585670602751%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3213876088517980551083924184682325205044405987565585670602751 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000000000000L'" (* 3213876088517980551083924184682325205044405987565585670602752 (0x200000000000000000000000000000000000000000000000000L) *)
- := (Const 3213876088517980551083924184682325205044405987565585670602752%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000L'" (* 3213876088517980551083924184682325205044405987565585670602752 (0x200000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000000000000001L'" (* 3213876088517980551083924184682325205044405987565585670602753 (0x200000000000000000000000000000000000000000000000001L) *)
- := (Const 3213876088517980551083924184682325205044405987565585670602753%Z).
-Notation "'0x200000000000000000000000000000000000000000000000001L'" (* 3213876088517980551083924184682325205044405987565585670602753 (0x200000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6427752177035961102167848369364650410088811975131171341205503 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 6427752177035961102167848369364650410088811975131171341205503%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6427752177035961102167848369364650410088811975131171341205503 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000000000000L'" (* 6427752177035961102167848369364650410088811975131171341205504 (0x400000000000000000000000000000000000000000000000000L) *)
- := (Const 6427752177035961102167848369364650410088811975131171341205504%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000L'" (* 6427752177035961102167848369364650410088811975131171341205504 (0x400000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000000000000001L'" (* 6427752177035961102167848369364650410088811975131171341205505 (0x400000000000000000000000000000000000000000000000001L) *)
- := (Const 6427752177035961102167848369364650410088811975131171341205505%Z).
-Notation "'0x400000000000000000000000000000000000000000000000001L'" (* 6427752177035961102167848369364650410088811975131171341205505 (0x400000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 12855504354071922204335696738729300820177623950262342682411007 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 12855504354071922204335696738729300820177623950262342682411007%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 12855504354071922204335696738729300820177623950262342682411007 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000000000000L'" (* 12855504354071922204335696738729300820177623950262342682411008 (0x800000000000000000000000000000000000000000000000000L) *)
- := (Const 12855504354071922204335696738729300820177623950262342682411008%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000L'" (* 12855504354071922204335696738729300820177623950262342682411008 (0x800000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000000000000001L'" (* 12855504354071922204335696738729300820177623950262342682411009 (0x800000000000000000000000000000000000000000000000001L) *)
- := (Const 12855504354071922204335696738729300820177623950262342682411009%Z).
-Notation "'0x800000000000000000000000000000000000000000000000001L'" (* 12855504354071922204335696738729300820177623950262342682411009 (0x800000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 25711008708143844408671393477458601640355247900524685364822015 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 25711008708143844408671393477458601640355247900524685364822015%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 25711008708143844408671393477458601640355247900524685364822015 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000000000000L'" (* 25711008708143844408671393477458601640355247900524685364822016 (0x1000000000000000000000000000000000000000000000000000L) *)
- := (Const 25711008708143844408671393477458601640355247900524685364822016%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000L'" (* 25711008708143844408671393477458601640355247900524685364822016 (0x1000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000000000000001L'" (* 25711008708143844408671393477458601640355247900524685364822017 (0x1000000000000000000000000000000000000000000000000001L) *)
- := (Const 25711008708143844408671393477458601640355247900524685364822017%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000001L'" (* 25711008708143844408671393477458601640355247900524685364822017 (0x1000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 51422017416287688817342786954917203280710495801049370729644031 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 51422017416287688817342786954917203280710495801049370729644031%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 51422017416287688817342786954917203280710495801049370729644031 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000000000000L'" (* 51422017416287688817342786954917203280710495801049370729644032 (0x2000000000000000000000000000000000000000000000000000L) *)
- := (Const 51422017416287688817342786954917203280710495801049370729644032%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000L'" (* 51422017416287688817342786954917203280710495801049370729644032 (0x2000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000000000000001L'" (* 51422017416287688817342786954917203280710495801049370729644033 (0x2000000000000000000000000000000000000000000000000001L) *)
- := (Const 51422017416287688817342786954917203280710495801049370729644033%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000001L'" (* 51422017416287688817342786954917203280710495801049370729644033 (0x2000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 102844034832575377634685573909834406561420991602098741459288063 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 102844034832575377634685573909834406561420991602098741459288063%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 102844034832575377634685573909834406561420991602098741459288063 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000000000000L'" (* 102844034832575377634685573909834406561420991602098741459288064 (0x4000000000000000000000000000000000000000000000000000L) *)
- := (Const 102844034832575377634685573909834406561420991602098741459288064%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000L'" (* 102844034832575377634685573909834406561420991602098741459288064 (0x4000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000000000000001L'" (* 102844034832575377634685573909834406561420991602098741459288065 (0x4000000000000000000000000000000000000000000000000001L) *)
- := (Const 102844034832575377634685573909834406561420991602098741459288065%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000001L'" (* 102844034832575377634685573909834406561420991602098741459288065 (0x4000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 205688069665150755269371147819668813122841983204197482918576127 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 205688069665150755269371147819668813122841983204197482918576127%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 205688069665150755269371147819668813122841983204197482918576127 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000000000000L'" (* 205688069665150755269371147819668813122841983204197482918576128 (0x8000000000000000000000000000000000000000000000000000L) *)
- := (Const 205688069665150755269371147819668813122841983204197482918576128%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000L'" (* 205688069665150755269371147819668813122841983204197482918576128 (0x8000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000000000000001L'" (* 205688069665150755269371147819668813122841983204197482918576129 (0x8000000000000000000000000000000000000000000000000001L) *)
- := (Const 205688069665150755269371147819668813122841983204197482918576129%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000001L'" (* 205688069665150755269371147819668813122841983204197482918576129 (0x8000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 411376139330301510538742295639337626245683966408394965837152255 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 411376139330301510538742295639337626245683966408394965837152255%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 411376139330301510538742295639337626245683966408394965837152255 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000000000000000L'" (* 411376139330301510538742295639337626245683966408394965837152256 (0x10000000000000000000000000000000000000000000000000000L) *)
- := (Const 411376139330301510538742295639337626245683966408394965837152256%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000000L'" (* 411376139330301510538742295639337626245683966408394965837152256 (0x10000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000000000000000001L'" (* 411376139330301510538742295639337626245683966408394965837152257 (0x10000000000000000000000000000000000000000000000000001L) *)
- := (Const 411376139330301510538742295639337626245683966408394965837152257%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000001L'" (* 411376139330301510538742295639337626245683966408394965837152257 (0x10000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 822752278660603021077484591278675252491367932816789931674304511 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 822752278660603021077484591278675252491367932816789931674304511%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 822752278660603021077484591278675252491367932816789931674304511 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000000000000000L'" (* 822752278660603021077484591278675252491367932816789931674304512 (0x20000000000000000000000000000000000000000000000000000L) *)
- := (Const 822752278660603021077484591278675252491367932816789931674304512%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000000L'" (* 822752278660603021077484591278675252491367932816789931674304512 (0x20000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000000000000000001L'" (* 822752278660603021077484591278675252491367932816789931674304513 (0x20000000000000000000000000000000000000000000000000001L) *)
- := (Const 822752278660603021077484591278675252491367932816789931674304513%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000001L'" (* 822752278660603021077484591278675252491367932816789931674304513 (0x20000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1645504557321206042154969182557350504982735865633579863348609023 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1645504557321206042154969182557350504982735865633579863348609023%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1645504557321206042154969182557350504982735865633579863348609023 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000000000000000L'" (* 1645504557321206042154969182557350504982735865633579863348609024 (0x40000000000000000000000000000000000000000000000000000L) *)
- := (Const 1645504557321206042154969182557350504982735865633579863348609024%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000000L'" (* 1645504557321206042154969182557350504982735865633579863348609024 (0x40000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000000000000000001L'" (* 1645504557321206042154969182557350504982735865633579863348609025 (0x40000000000000000000000000000000000000000000000000001L) *)
- := (Const 1645504557321206042154969182557350504982735865633579863348609025%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000001L'" (* 1645504557321206042154969182557350504982735865633579863348609025 (0x40000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3291009114642412084309938365114701009965471731267159726697218047 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3291009114642412084309938365114701009965471731267159726697218047%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3291009114642412084309938365114701009965471731267159726697218047 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000000000000000L'" (* 3291009114642412084309938365114701009965471731267159726697218048 (0x80000000000000000000000000000000000000000000000000000L) *)
- := (Const 3291009114642412084309938365114701009965471731267159726697218048%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000000L'" (* 3291009114642412084309938365114701009965471731267159726697218048 (0x80000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000000000000000001L'" (* 3291009114642412084309938365114701009965471731267159726697218049 (0x80000000000000000000000000000000000000000000000000001L) *)
- := (Const 3291009114642412084309938365114701009965471731267159726697218049%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000001L'" (* 3291009114642412084309938365114701009965471731267159726697218049 (0x80000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6582018229284824168619876730229402019930943462534319453394436095 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 6582018229284824168619876730229402019930943462534319453394436095%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6582018229284824168619876730229402019930943462534319453394436095 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000000000000000L'" (* 6582018229284824168619876730229402019930943462534319453394436096 (0x100000000000000000000000000000000000000000000000000000L) *)
- := (Const 6582018229284824168619876730229402019930943462534319453394436096%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000000L'" (* 6582018229284824168619876730229402019930943462534319453394436096 (0x100000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000000000000000001L'" (* 6582018229284824168619876730229402019930943462534319453394436097 (0x100000000000000000000000000000000000000000000000000001L) *)
- := (Const 6582018229284824168619876730229402019930943462534319453394436097%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000001L'" (* 6582018229284824168619876730229402019930943462534319453394436097 (0x100000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 13164036458569648337239753460458804039861886925068638906788872191 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 13164036458569648337239753460458804039861886925068638906788872191%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 13164036458569648337239753460458804039861886925068638906788872191 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000000000000000L'" (* 13164036458569648337239753460458804039861886925068638906788872192 (0x200000000000000000000000000000000000000000000000000000L) *)
- := (Const 13164036458569648337239753460458804039861886925068638906788872192%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000000L'" (* 13164036458569648337239753460458804039861886925068638906788872192 (0x200000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000000000000000001L'" (* 13164036458569648337239753460458804039861886925068638906788872193 (0x200000000000000000000000000000000000000000000000000001L) *)
- := (Const 13164036458569648337239753460458804039861886925068638906788872193%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000001L'" (* 13164036458569648337239753460458804039861886925068638906788872193 (0x200000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 26328072917139296674479506920917608079723773850137277813577744383 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 26328072917139296674479506920917608079723773850137277813577744383%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 26328072917139296674479506920917608079723773850137277813577744383 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000000000000000L'" (* 26328072917139296674479506920917608079723773850137277813577744384 (0x400000000000000000000000000000000000000000000000000000L) *)
- := (Const 26328072917139296674479506920917608079723773850137277813577744384%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000000L'" (* 26328072917139296674479506920917608079723773850137277813577744384 (0x400000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000000000000000001L'" (* 26328072917139296674479506920917608079723773850137277813577744385 (0x400000000000000000000000000000000000000000000000000001L) *)
- := (Const 26328072917139296674479506920917608079723773850137277813577744385%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000001L'" (* 26328072917139296674479506920917608079723773850137277813577744385 (0x400000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 52656145834278593348959013841835216159447547700274555627155488767 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 52656145834278593348959013841835216159447547700274555627155488767%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 52656145834278593348959013841835216159447547700274555627155488767 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000000000000000L'" (* 52656145834278593348959013841835216159447547700274555627155488768 (0x800000000000000000000000000000000000000000000000000000L) *)
- := (Const 52656145834278593348959013841835216159447547700274555627155488768%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000000L'" (* 52656145834278593348959013841835216159447547700274555627155488768 (0x800000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000000000000000001L'" (* 52656145834278593348959013841835216159447547700274555627155488769 (0x800000000000000000000000000000000000000000000000000001L) *)
- := (Const 52656145834278593348959013841835216159447547700274555627155488769%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000001L'" (* 52656145834278593348959013841835216159447547700274555627155488769 (0x800000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 105312291668557186697918027683670432318895095400549111254310977535 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 105312291668557186697918027683670432318895095400549111254310977535%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 105312291668557186697918027683670432318895095400549111254310977535 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000000000000000L'" (* 105312291668557186697918027683670432318895095400549111254310977536 (0x1000000000000000000000000000000000000000000000000000000L) *)
- := (Const 105312291668557186697918027683670432318895095400549111254310977536%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000000L'" (* 105312291668557186697918027683670432318895095400549111254310977536 (0x1000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000000000000000001L'" (* 105312291668557186697918027683670432318895095400549111254310977537 (0x1000000000000000000000000000000000000000000000000000001L) *)
- := (Const 105312291668557186697918027683670432318895095400549111254310977537%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000001L'" (* 105312291668557186697918027683670432318895095400549111254310977537 (0x1000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 210624583337114373395836055367340864637790190801098222508621955071 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 210624583337114373395836055367340864637790190801098222508621955071%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 210624583337114373395836055367340864637790190801098222508621955071 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000000000000000L'" (* 210624583337114373395836055367340864637790190801098222508621955072 (0x2000000000000000000000000000000000000000000000000000000L) *)
- := (Const 210624583337114373395836055367340864637790190801098222508621955072%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000000L'" (* 210624583337114373395836055367340864637790190801098222508621955072 (0x2000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000000000000000001L'" (* 210624583337114373395836055367340864637790190801098222508621955073 (0x2000000000000000000000000000000000000000000000000000001L) *)
- := (Const 210624583337114373395836055367340864637790190801098222508621955073%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000001L'" (* 210624583337114373395836055367340864637790190801098222508621955073 (0x2000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 421249166674228746791672110734681729275580381602196445017243910143 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 421249166674228746791672110734681729275580381602196445017243910143%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 421249166674228746791672110734681729275580381602196445017243910143 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000000000000000L'" (* 421249166674228746791672110734681729275580381602196445017243910144 (0x4000000000000000000000000000000000000000000000000000000L) *)
- := (Const 421249166674228746791672110734681729275580381602196445017243910144%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000000L'" (* 421249166674228746791672110734681729275580381602196445017243910144 (0x4000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000000000000000001L'" (* 421249166674228746791672110734681729275580381602196445017243910145 (0x4000000000000000000000000000000000000000000000000000001L) *)
- := (Const 421249166674228746791672110734681729275580381602196445017243910145%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000001L'" (* 421249166674228746791672110734681729275580381602196445017243910145 (0x4000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 842498333348457493583344221469363458551160763204392890034487820287 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 842498333348457493583344221469363458551160763204392890034487820287%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 842498333348457493583344221469363458551160763204392890034487820287 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000000000000000L'" (* 842498333348457493583344221469363458551160763204392890034487820288 (0x8000000000000000000000000000000000000000000000000000000L) *)
- := (Const 842498333348457493583344221469363458551160763204392890034487820288%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000000L'" (* 842498333348457493583344221469363458551160763204392890034487820288 (0x8000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000000000000000001L'" (* 842498333348457493583344221469363458551160763204392890034487820289 (0x8000000000000000000000000000000000000000000000000000001L) *)
- := (Const 842498333348457493583344221469363458551160763204392890034487820289%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000001L'" (* 842498333348457493583344221469363458551160763204392890034487820289 (0x8000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1684996666696914987166688442938726917102321526408785780068975640575 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1684996666696914987166688442938726917102321526408785780068975640575%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1684996666696914987166688442938726917102321526408785780068975640575 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000000000000000000L'" (* 1684996666696914987166688442938726917102321526408785780068975640576 (0x10000000000000000000000000000000000000000000000000000000L) *)
- := (Const 1684996666696914987166688442938726917102321526408785780068975640576%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000000000L'" (* 1684996666696914987166688442938726917102321526408785780068975640576 (0x10000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000000000000000000001L'" (* 1684996666696914987166688442938726917102321526408785780068975640577 (0x10000000000000000000000000000000000000000000000000000001L) *)
- := (Const 1684996666696914987166688442938726917102321526408785780068975640577%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000000001L'" (* 1684996666696914987166688442938726917102321526408785780068975640577 (0x10000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3369993333393829974333376885877453834204643052817571560137951281151 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3369993333393829974333376885877453834204643052817571560137951281151%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3369993333393829974333376885877453834204643052817571560137951281151 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000000000000000000L'" (* 3369993333393829974333376885877453834204643052817571560137951281152 (0x20000000000000000000000000000000000000000000000000000000L) *)
- := (Const 3369993333393829974333376885877453834204643052817571560137951281152%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000000000L'" (* 3369993333393829974333376885877453834204643052817571560137951281152 (0x20000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000000000000000000001L'" (* 3369993333393829974333376885877453834204643052817571560137951281153 (0x20000000000000000000000000000000000000000000000000000001L) *)
- := (Const 3369993333393829974333376885877453834204643052817571560137951281153%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000000001L'" (* 3369993333393829974333376885877453834204643052817571560137951281153 (0x20000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6739986666787659948666753771754907668409286105635143120275902562303 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 6739986666787659948666753771754907668409286105635143120275902562303%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6739986666787659948666753771754907668409286105635143120275902562303 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000000000000000000L'" (* 6739986666787659948666753771754907668409286105635143120275902562304 (0x40000000000000000000000000000000000000000000000000000000L) *)
- := (Const 6739986666787659948666753771754907668409286105635143120275902562304%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000000000L'" (* 6739986666787659948666753771754907668409286105635143120275902562304 (0x40000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000000000000000000001L'" (* 6739986666787659948666753771754907668409286105635143120275902562305 (0x40000000000000000000000000000000000000000000000000000001L) *)
- := (Const 6739986666787659948666753771754907668409286105635143120275902562305%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000000001L'" (* 6739986666787659948666753771754907668409286105635143120275902562305 (0x40000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 13479973333575319897333507543509815336818572211270286240551805124607 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 13479973333575319897333507543509815336818572211270286240551805124607%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 13479973333575319897333507543509815336818572211270286240551805124607 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000000000000000000L'" (* 13479973333575319897333507543509815336818572211270286240551805124608 (0x80000000000000000000000000000000000000000000000000000000L) *)
- := (Const 13479973333575319897333507543509815336818572211270286240551805124608%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000000000L'" (* 13479973333575319897333507543509815336818572211270286240551805124608 (0x80000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000000000000000000001L'" (* 13479973333575319897333507543509815336818572211270286240551805124609 (0x80000000000000000000000000000000000000000000000000000001L) *)
- := (Const 13479973333575319897333507543509815336818572211270286240551805124609%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000000001L'" (* 13479973333575319897333507543509815336818572211270286240551805124609 (0x80000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 26959946667150639794667015087019630673637144422540572481103610249215 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 26959946667150639794667015087019630673637144422540572481103610249215%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 26959946667150639794667015087019630673637144422540572481103610249215 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000000000000000000L'" (* 26959946667150639794667015087019630673637144422540572481103610249216 (0x100000000000000000000000000000000000000000000000000000000L) *)
- := (Const 26959946667150639794667015087019630673637144422540572481103610249216%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000000000L'" (* 26959946667150639794667015087019630673637144422540572481103610249216 (0x100000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000000000000000000001L'" (* 26959946667150639794667015087019630673637144422540572481103610249217 (0x100000000000000000000000000000000000000000000000000000001L) *)
- := (Const 26959946667150639794667015087019630673637144422540572481103610249217%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000000001L'" (* 26959946667150639794667015087019630673637144422540572481103610249217 (0x100000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 53919893334301279589334030174039261347274288845081144962207220498431 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 53919893334301279589334030174039261347274288845081144962207220498431%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 53919893334301279589334030174039261347274288845081144962207220498431 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000000000000000000L'" (* 53919893334301279589334030174039261347274288845081144962207220498432 (0x200000000000000000000000000000000000000000000000000000000L) *)
- := (Const 53919893334301279589334030174039261347274288845081144962207220498432%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000000000L'" (* 53919893334301279589334030174039261347274288845081144962207220498432 (0x200000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000000000000000000001L'" (* 53919893334301279589334030174039261347274288845081144962207220498433 (0x200000000000000000000000000000000000000000000000000000001L) *)
- := (Const 53919893334301279589334030174039261347274288845081144962207220498433%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000000001L'" (* 53919893334301279589334030174039261347274288845081144962207220498433 (0x200000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 107839786668602559178668060348078522694548577690162289924414440996863 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 107839786668602559178668060348078522694548577690162289924414440996863%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 107839786668602559178668060348078522694548577690162289924414440996863 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000000000000000000L'" (* 107839786668602559178668060348078522694548577690162289924414440996864 (0x400000000000000000000000000000000000000000000000000000000L) *)
- := (Const 107839786668602559178668060348078522694548577690162289924414440996864%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000000000L'" (* 107839786668602559178668060348078522694548577690162289924414440996864 (0x400000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000000000000000000001L'" (* 107839786668602559178668060348078522694548577690162289924414440996865 (0x400000000000000000000000000000000000000000000000000000001L) *)
- := (Const 107839786668602559178668060348078522694548577690162289924414440996865%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000000001L'" (* 107839786668602559178668060348078522694548577690162289924414440996865 (0x400000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 215679573337205118357336120696157045389097155380324579848828881993727 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 215679573337205118357336120696157045389097155380324579848828881993727%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 215679573337205118357336120696157045389097155380324579848828881993727 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000000000000000000L'" (* 215679573337205118357336120696157045389097155380324579848828881993728 (0x800000000000000000000000000000000000000000000000000000000L) *)
- := (Const 215679573337205118357336120696157045389097155380324579848828881993728%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000000000L'" (* 215679573337205118357336120696157045389097155380324579848828881993728 (0x800000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000000000000000000001L'" (* 215679573337205118357336120696157045389097155380324579848828881993729 (0x800000000000000000000000000000000000000000000000000000001L) *)
- := (Const 215679573337205118357336120696157045389097155380324579848828881993729%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000000001L'" (* 215679573337205118357336120696157045389097155380324579848828881993729 (0x800000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 431359146674410236714672241392314090778194310760649159697657763987455 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 431359146674410236714672241392314090778194310760649159697657763987455%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 431359146674410236714672241392314090778194310760649159697657763987455 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000000000000000000L'" (* 431359146674410236714672241392314090778194310760649159697657763987456 (0x1000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 431359146674410236714672241392314090778194310760649159697657763987456%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000000000L'" (* 431359146674410236714672241392314090778194310760649159697657763987456 (0x1000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000000000000000000001L'" (* 431359146674410236714672241392314090778194310760649159697657763987457 (0x1000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 431359146674410236714672241392314090778194310760649159697657763987457%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000000001L'" (* 431359146674410236714672241392314090778194310760649159697657763987457 (0x1000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 862718293348820473429344482784628181556388621521298319395315527974911 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 862718293348820473429344482784628181556388621521298319395315527974911%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 862718293348820473429344482784628181556388621521298319395315527974911 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000000000000000000L'" (* 862718293348820473429344482784628181556388621521298319395315527974912 (0x2000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 862718293348820473429344482784628181556388621521298319395315527974912%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000000000L'" (* 862718293348820473429344482784628181556388621521298319395315527974912 (0x2000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000000000000000000001L'" (* 862718293348820473429344482784628181556388621521298319395315527974913 (0x2000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 862718293348820473429344482784628181556388621521298319395315527974913%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000000001L'" (* 862718293348820473429344482784628181556388621521298319395315527974913 (0x2000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1725436586697640946858688965569256363112777243042596638790631055949823 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1725436586697640946858688965569256363112777243042596638790631055949823%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1725436586697640946858688965569256363112777243042596638790631055949823 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000000000000000000L'" (* 1725436586697640946858688965569256363112777243042596638790631055949824 (0x4000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 1725436586697640946858688965569256363112777243042596638790631055949824%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000000000L'" (* 1725436586697640946858688965569256363112777243042596638790631055949824 (0x4000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000000000000000000001L'" (* 1725436586697640946858688965569256363112777243042596638790631055949825 (0x4000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 1725436586697640946858688965569256363112777243042596638790631055949825%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000000001L'" (* 1725436586697640946858688965569256363112777243042596638790631055949825 (0x4000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3450873173395281893717377931138512726225554486085193277581262111899647 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3450873173395281893717377931138512726225554486085193277581262111899647%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3450873173395281893717377931138512726225554486085193277581262111899647 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000000000000000000L'" (* 3450873173395281893717377931138512726225554486085193277581262111899648 (0x8000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 3450873173395281893717377931138512726225554486085193277581262111899648%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000000000L'" (* 3450873173395281893717377931138512726225554486085193277581262111899648 (0x8000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000000000000000000001L'" (* 3450873173395281893717377931138512726225554486085193277581262111899649 (0x8000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 3450873173395281893717377931138512726225554486085193277581262111899649%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000000001L'" (* 3450873173395281893717377931138512726225554486085193277581262111899649 (0x8000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6901746346790563787434755862277025452451108972170386555162524223799295 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 6901746346790563787434755862277025452451108972170386555162524223799295%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 6901746346790563787434755862277025452451108972170386555162524223799295 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000000000000000000000L'" (* 6901746346790563787434755862277025452451108972170386555162524223799296 (0x10000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 6901746346790563787434755862277025452451108972170386555162524223799296%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000000000000L'" (* 6901746346790563787434755862277025452451108972170386555162524223799296 (0x10000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000000000000000000000001L'" (* 6901746346790563787434755862277025452451108972170386555162524223799297 (0x10000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 6901746346790563787434755862277025452451108972170386555162524223799297%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000000000001L'" (* 6901746346790563787434755862277025452451108972170386555162524223799297 (0x10000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 13803492693581127574869511724554050904902217944340773110325048447598591 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 13803492693581127574869511724554050904902217944340773110325048447598591%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 13803492693581127574869511724554050904902217944340773110325048447598591 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000000000000000000000L'" (* 13803492693581127574869511724554050904902217944340773110325048447598592 (0x20000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 13803492693581127574869511724554050904902217944340773110325048447598592%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000000000000L'" (* 13803492693581127574869511724554050904902217944340773110325048447598592 (0x20000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000000000000000000000001L'" (* 13803492693581127574869511724554050904902217944340773110325048447598593 (0x20000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 13803492693581127574869511724554050904902217944340773110325048447598593%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000000000001L'" (* 13803492693581127574869511724554050904902217944340773110325048447598593 (0x20000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 27606985387162255149739023449108101809804435888681546220650096895197183 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 27606985387162255149739023449108101809804435888681546220650096895197183%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 27606985387162255149739023449108101809804435888681546220650096895197183 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000000000000000000000L'" (* 27606985387162255149739023449108101809804435888681546220650096895197184 (0x40000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 27606985387162255149739023449108101809804435888681546220650096895197184%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000000000000L'" (* 27606985387162255149739023449108101809804435888681546220650096895197184 (0x40000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000000000000000000000001L'" (* 27606985387162255149739023449108101809804435888681546220650096895197185 (0x40000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 27606985387162255149739023449108101809804435888681546220650096895197185%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000000000001L'" (* 27606985387162255149739023449108101809804435888681546220650096895197185 (0x40000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 55213970774324510299478046898216203619608871777363092441300193790394367 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 55213970774324510299478046898216203619608871777363092441300193790394367%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 55213970774324510299478046898216203619608871777363092441300193790394367 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000000000000000000000L'" (* 55213970774324510299478046898216203619608871777363092441300193790394368 (0x80000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 55213970774324510299478046898216203619608871777363092441300193790394368%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000000000000L'" (* 55213970774324510299478046898216203619608871777363092441300193790394368 (0x80000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000000000000000000000001L'" (* 55213970774324510299478046898216203619608871777363092441300193790394369 (0x80000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 55213970774324510299478046898216203619608871777363092441300193790394369%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000000000001L'" (* 55213970774324510299478046898216203619608871777363092441300193790394369 (0x80000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 110427941548649020598956093796432407239217743554726184882600387580788735 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 110427941548649020598956093796432407239217743554726184882600387580788735%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 110427941548649020598956093796432407239217743554726184882600387580788735 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000000000000000000000L'" (* 110427941548649020598956093796432407239217743554726184882600387580788736 (0x100000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 110427941548649020598956093796432407239217743554726184882600387580788736%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000000000000L'" (* 110427941548649020598956093796432407239217743554726184882600387580788736 (0x100000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000000000000000000000001L'" (* 110427941548649020598956093796432407239217743554726184882600387580788737 (0x100000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 110427941548649020598956093796432407239217743554726184882600387580788737%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000000000001L'" (* 110427941548649020598956093796432407239217743554726184882600387580788737 (0x100000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 220855883097298041197912187592864814478435487109452369765200775161577471 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 220855883097298041197912187592864814478435487109452369765200775161577471%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 220855883097298041197912187592864814478435487109452369765200775161577471 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000000000000000000000L'" (* 220855883097298041197912187592864814478435487109452369765200775161577472 (0x200000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 220855883097298041197912187592864814478435487109452369765200775161577472%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000000000000L'" (* 220855883097298041197912187592864814478435487109452369765200775161577472 (0x200000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000000000000000000000001L'" (* 220855883097298041197912187592864814478435487109452369765200775161577473 (0x200000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 220855883097298041197912187592864814478435487109452369765200775161577473%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000000000001L'" (* 220855883097298041197912187592864814478435487109452369765200775161577473 (0x200000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 441711766194596082395824375185729628956870974218904739530401550323154943 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 441711766194596082395824375185729628956870974218904739530401550323154943%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 441711766194596082395824375185729628956870974218904739530401550323154943 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000000000000000000000L'" (* 441711766194596082395824375185729628956870974218904739530401550323154944 (0x400000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 441711766194596082395824375185729628956870974218904739530401550323154944%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000000000000L'" (* 441711766194596082395824375185729628956870974218904739530401550323154944 (0x400000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000000000000000000000001L'" (* 441711766194596082395824375185729628956870974218904739530401550323154945 (0x400000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 441711766194596082395824375185729628956870974218904739530401550323154945%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000000000001L'" (* 441711766194596082395824375185729628956870974218904739530401550323154945 (0x400000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 883423532389192164791648750371459257913741948437809479060803100646309887 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 883423532389192164791648750371459257913741948437809479060803100646309887%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 883423532389192164791648750371459257913741948437809479060803100646309887 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000000000000000000000L'" (* 883423532389192164791648750371459257913741948437809479060803100646309888 (0x800000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 883423532389192164791648750371459257913741948437809479060803100646309888%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000000000000L'" (* 883423532389192164791648750371459257913741948437809479060803100646309888 (0x800000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000000000000000000000001L'" (* 883423532389192164791648750371459257913741948437809479060803100646309889 (0x800000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 883423532389192164791648750371459257913741948437809479060803100646309889%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000000000001L'" (* 883423532389192164791648750371459257913741948437809479060803100646309889 (0x800000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1766847064778384329583297500742918515827483896875618958121606201292619775 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1766847064778384329583297500742918515827483896875618958121606201292619775%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1766847064778384329583297500742918515827483896875618958121606201292619775 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000000000000000000000L'" (* 1766847064778384329583297500742918515827483896875618958121606201292619776 (0x1000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 1766847064778384329583297500742918515827483896875618958121606201292619776%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000000000000L'" (* 1766847064778384329583297500742918515827483896875618958121606201292619776 (0x1000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000000000000000000000001L'" (* 1766847064778384329583297500742918515827483896875618958121606201292619777 (0x1000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 1766847064778384329583297500742918515827483896875618958121606201292619777%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000000000001L'" (* 1766847064778384329583297500742918515827483896875618958121606201292619777 (0x1000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3533694129556768659166595001485837031654967793751237916243212402585239551 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3533694129556768659166595001485837031654967793751237916243212402585239551%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3533694129556768659166595001485837031654967793751237916243212402585239551 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000000000000000000000L'" (* 3533694129556768659166595001485837031654967793751237916243212402585239552 (0x2000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 3533694129556768659166595001485837031654967793751237916243212402585239552%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000000000000L'" (* 3533694129556768659166595001485837031654967793751237916243212402585239552 (0x2000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000000000000000000000001L'" (* 3533694129556768659166595001485837031654967793751237916243212402585239553 (0x2000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 3533694129556768659166595001485837031654967793751237916243212402585239553%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000000000001L'" (* 3533694129556768659166595001485837031654967793751237916243212402585239553 (0x2000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 7067388259113537318333190002971674063309935587502475832486424805170479103 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 7067388259113537318333190002971674063309935587502475832486424805170479103%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 7067388259113537318333190002971674063309935587502475832486424805170479103 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~0~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000000000000000000000L'" (* 7067388259113537318333190002971674063309935587502475832486424805170479104 (0x4000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 7067388259113537318333190002971674063309935587502475832486424805170479104%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000000000000L'" (* 7067388259113537318333190002971674063309935587502475832486424805170479104 (0x4000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000000000000000000000001L'" (* 7067388259113537318333190002971674063309935587502475832486424805170479105 (0x4000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 7067388259113537318333190002971674063309935587502475832486424805170479105%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000000000001L'" (* 7067388259113537318333190002971674063309935587502475832486424805170479105 (0x4000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 14134776518227074636666380005943348126619871175004951664972849610340958207 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 14134776518227074636666380005943348126619871175004951664972849610340958207%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 14134776518227074636666380005943348126619871175004951664972849610340958207 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000000000000000000000L'" (* 14134776518227074636666380005943348126619871175004951664972849610340958208 (0x8000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 14134776518227074636666380005943348126619871175004951664972849610340958208%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000000000000L'" (* 14134776518227074636666380005943348126619871175004951664972849610340958208 (0x8000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000000000000000000000001L'" (* 14134776518227074636666380005943348126619871175004951664972849610340958209 (0x8000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 14134776518227074636666380005943348126619871175004951664972849610340958209%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000000000001L'" (* 14134776518227074636666380005943348126619871175004951664972849610340958209 (0x8000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 28269553036454149273332760011886696253239742350009903329945699220681916415 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 28269553036454149273332760011886696253239742350009903329945699220681916415%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 28269553036454149273332760011886696253239742350009903329945699220681916415 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x10000000000000000000000000000000000000000000000000000000000000L'" (* 28269553036454149273332760011886696253239742350009903329945699220681916416 (0x10000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 28269553036454149273332760011886696253239742350009903329945699220681916416%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000000000000000L'" (* 28269553036454149273332760011886696253239742350009903329945699220681916416 (0x10000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x10000000000000000000000000000000000000000000000000000000000001L'" (* 28269553036454149273332760011886696253239742350009903329945699220681916417 (0x10000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 28269553036454149273332760011886696253239742350009903329945699220681916417%Z).
-Notation "'0x10000000000000000000000000000000000000000000000000000000000001L'" (* 28269553036454149273332760011886696253239742350009903329945699220681916417 (0x10000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 56539106072908298546665520023773392506479484700019806659891398441363832831 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 56539106072908298546665520023773392506479484700019806659891398441363832831%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 56539106072908298546665520023773392506479484700019806659891398441363832831 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x20000000000000000000000000000000000000000000000000000000000000L'" (* 56539106072908298546665520023773392506479484700019806659891398441363832832 (0x20000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 56539106072908298546665520023773392506479484700019806659891398441363832832%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000000000000000L'" (* 56539106072908298546665520023773392506479484700019806659891398441363832832 (0x20000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x20000000000000000000000000000000000000000000000000000000000001L'" (* 56539106072908298546665520023773392506479484700019806659891398441363832833 (0x20000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 56539106072908298546665520023773392506479484700019806659891398441363832833%Z).
-Notation "'0x20000000000000000000000000000000000000000000000000000000000001L'" (* 56539106072908298546665520023773392506479484700019806659891398441363832833 (0x20000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 113078212145816597093331040047546785012958969400039613319782796882727665663 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 113078212145816597093331040047546785012958969400039613319782796882727665663%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 113078212145816597093331040047546785012958969400039613319782796882727665663 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x40000000000000000000000000000000000000000000000000000000000000L'" (* 113078212145816597093331040047546785012958969400039613319782796882727665664 (0x40000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 113078212145816597093331040047546785012958969400039613319782796882727665664%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000000000000000L'" (* 113078212145816597093331040047546785012958969400039613319782796882727665664 (0x40000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x40000000000000000000000000000000000000000000000000000000000001L'" (* 113078212145816597093331040047546785012958969400039613319782796882727665665 (0x40000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 113078212145816597093331040047546785012958969400039613319782796882727665665%Z).
-Notation "'0x40000000000000000000000000000000000000000000000000000000000001L'" (* 113078212145816597093331040047546785012958969400039613319782796882727665665 (0x40000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 226156424291633194186662080095093570025917938800079226639565593765455331327 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 226156424291633194186662080095093570025917938800079226639565593765455331327%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 226156424291633194186662080095093570025917938800079226639565593765455331327 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x80000000000000000000000000000000000000000000000000000000000000L'" (* 226156424291633194186662080095093570025917938800079226639565593765455331328 (0x80000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 226156424291633194186662080095093570025917938800079226639565593765455331328%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000000000000000L'" (* 226156424291633194186662080095093570025917938800079226639565593765455331328 (0x80000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x80000000000000000000000000000000000000000000000000000000000001L'" (* 226156424291633194186662080095093570025917938800079226639565593765455331329 (0x80000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 226156424291633194186662080095093570025917938800079226639565593765455331329%Z).
-Notation "'0x80000000000000000000000000000000000000000000000000000000000001L'" (* 226156424291633194186662080095093570025917938800079226639565593765455331329 (0x80000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 452312848583266388373324160190187140051835877600158453279131187530910662655 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 452312848583266388373324160190187140051835877600158453279131187530910662655%Z).
-Notation "'0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 452312848583266388373324160190187140051835877600158453279131187530910662655 (0xffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x100000000000000000000000000000000000000000000000000000000000000L'" (* 452312848583266388373324160190187140051835877600158453279131187530910662656 (0x100000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 452312848583266388373324160190187140051835877600158453279131187530910662656%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000000000000000L'" (* 452312848583266388373324160190187140051835877600158453279131187530910662656 (0x100000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x100000000000000000000000000000000000000000000000000000000000001L'" (* 452312848583266388373324160190187140051835877600158453279131187530910662657 (0x100000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 452312848583266388373324160190187140051835877600158453279131187530910662657%Z).
-Notation "'0x100000000000000000000000000000000000000000000000000000000000001L'" (* 452312848583266388373324160190187140051835877600158453279131187530910662657 (0x100000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 904625697166532776746648320380374280103671755200316906558262375061821325311 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 904625697166532776746648320380374280103671755200316906558262375061821325311%Z).
-Notation "'0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 904625697166532776746648320380374280103671755200316906558262375061821325311 (0x1ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x200000000000000000000000000000000000000000000000000000000000000L'" (* 904625697166532776746648320380374280103671755200316906558262375061821325312 (0x200000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 904625697166532776746648320380374280103671755200316906558262375061821325312%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000000000000000L'" (* 904625697166532776746648320380374280103671755200316906558262375061821325312 (0x200000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x200000000000000000000000000000000000000000000000000000000000001L'" (* 904625697166532776746648320380374280103671755200316906558262375061821325313 (0x200000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 904625697166532776746648320380374280103671755200316906558262375061821325313%Z).
-Notation "'0x200000000000000000000000000000000000000000000000000000000000001L'" (* 904625697166532776746648320380374280103671755200316906558262375061821325313 (0x200000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650623 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 1809251394333065553493296640760748560207343510400633813116524750123642650623%Z).
-Notation "'0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650623 (0x3ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x400000000000000000000000000000000000000000000000000000000000000L'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650624 (0x400000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 1809251394333065553493296640760748560207343510400633813116524750123642650624%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000000000000000L'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650624 (0x400000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x400000000000000000000000000000000000000000000000000000000000001L'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650625 (0x400000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 1809251394333065553493296640760748560207343510400633813116524750123642650625%Z).
-Notation "'0x400000000000000000000000000000000000000000000000000000000000001L'" (* 1809251394333065553493296640760748560207343510400633813116524750123642650625 (0x400000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301247 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 3618502788666131106986593281521497120414687020801267626233049500247285301247%Z).
-Notation "'0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301247 (0x7ffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x800000000000000000000000000000000000000000000000000000000000000L'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301248 (0x800000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 3618502788666131106986593281521497120414687020801267626233049500247285301248%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000000000000000L'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301248 (0x800000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x800000000000000000000000000000000000000000000000000000000000001L'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301249 (0x800000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 3618502788666131106986593281521497120414687020801267626233049500247285301249%Z).
-Notation "'0x800000000000000000000000000000000000000000000000000000000000001L'" (* 3618502788666131106986593281521497120414687020801267626233049500247285301249 (0x800000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602495 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 7237005577332262213973186563042994240829374041602535252466099000494570602495%Z).
-Notation "'0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602495 (0xfffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x1000000000000000000000000000000000000000000000000000000000000000L'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602496 (0x1000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 7237005577332262213973186563042994240829374041602535252466099000494570602496%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000000000000000L'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602496 (0x1000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x1000000000000000000000000000000000000000000000000000000000000001L'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602497 (0x1000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 7237005577332262213973186563042994240829374041602535252466099000494570602497%Z).
-Notation "'0x1000000000000000000000000000000000000000000000000000000000000001L'" (* 7237005577332262213973186563042994240829374041602535252466099000494570602497 (0x1000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204991 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 14474011154664524427946373126085988481658748083205070504932198000989141204991%Z).
-Notation "'0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204991 (0x1fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x2000000000000000000000000000000000000000000000000000000000000000L'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204992 (0x2000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 14474011154664524427946373126085988481658748083205070504932198000989141204992%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000000000000000L'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204992 (0x2000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x2000000000000000000000000000000000000000000000000000000000000001L'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204993 (0x2000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 14474011154664524427946373126085988481658748083205070504932198000989141204993%Z).
-Notation "'0x2000000000000000000000000000000000000000000000000000000000000001L'" (* 14474011154664524427946373126085988481658748083205070504932198000989141204993 (0x2000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409983 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 28948022309329048855892746252171976963317496166410141009864396001978282409983%Z).
-Notation "'0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409983 (0x3fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x4000000000000000000000000000000000000000000000000000000000000000L'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409984 (0x4000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 28948022309329048855892746252171976963317496166410141009864396001978282409984%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000000000000000L'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409984 (0x4000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x4000000000000000000000000000000000000000000000000000000000000001L'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409985 (0x4000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 28948022309329048855892746252171976963317496166410141009864396001978282409985%Z).
-Notation "'0x4000000000000000000000000000000000000000000000000000000000000001L'" (* 28948022309329048855892746252171976963317496166410141009864396001978282409985 (0x4000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~0~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819967 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const 57896044618658097711785492504343953926634992332820282019728792003956564819967%Z).
-Notation "'0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819967 (0x7fffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffffL) *)
- := (Const WO~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~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~1~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 "'0x8000000000000000000000000000000000000000000000000000000000000000L'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819968 (0x8000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const 57896044618658097711785492504343953926634992332820282019728792003956564819968%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000000000000000L'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819968 (0x8000000000000000000000000000000000000000000000000000000000000000L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0).
-Notation "'0x8000000000000000000000000000000000000000000000000000000000000001L'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819969 (0x8000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const 57896044618658097711785492504343953926634992332820282019728792003956564819969%Z).
-Notation "'0x8000000000000000000000000000000000000000000000000000000000000001L'" (* 57896044618658097711785492504343953926634992332820282019728792003956564819969 (0x8000000000000000000000000000000000000000000000000000000000000001L) *)
- := (Const WO~1~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~0~1).
diff --git a/src/Compilers/Z/Inline.v b/src/Compilers/Z/Inline.v
deleted file mode 100644
index cbe6dbfde..000000000
--- a/src/Compilers/Z/Inline.v
+++ /dev/null
@@ -1,10 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Inline.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Util.
-
-Definition InlineConstAndOpp {t} (e : Expr t) : Expr t
- := @InlineConst base_type op (is_const_or_opp) t e.
-
-Definition InlineConst {t} (e : Expr t) : Expr t
- := @InlineConst base_type op (is_const) t e.
diff --git a/src/Compilers/Z/InlineConstAndOp.v b/src/Compilers/Z/InlineConstAndOp.v
deleted file mode 100644
index ad2a2ab58..000000000
--- a/src/Compilers/Z/InlineConstAndOp.v
+++ /dev/null
@@ -1,11 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InlineConstAndOp.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Util.
-
-Definition inline_const_and_opf {var} {t} (e : exprf _ _ t) : @exprf base_type op var t
- := @inline_const_and_opf base_type op interp_base_type (@interp_op) var make_const t e.
-Definition inline_const_and_op {var} {t} (e : expr _ _ t) : @expr base_type op var t
- := @inline_const_and_op base_type op interp_base_type (@interp_op) var make_const t e.
-Definition InlineConstAndOp {t} (e : Expr t) : Expr t
- := @InlineConstAndOp base_type op interp_base_type interp_op make_const t e.
diff --git a/src/Compilers/Z/InlineConstAndOpByRewrite.v b/src/Compilers/Z/InlineConstAndOpByRewrite.v
deleted file mode 100644
index f4403e12b..000000000
--- a/src/Compilers/Z/InlineConstAndOpByRewrite.v
+++ /dev/null
@@ -1,13 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InlineConstAndOpByRewrite.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Util.
-
-Module Export Rewrite.
- Definition inline_const_and_opf {var} {t} (e : exprf _ _ t) : @exprf base_type op var t
- := @inline_const_and_opf base_type op interp_base_type (@interp_op) var make_const t e.
- Definition inline_const_and_op {var} {t} (e : expr _ _ t) : @expr base_type op var t
- := @inline_const_and_op base_type op interp_base_type (@interp_op) var make_const t e.
- Definition InlineConstAndOp {t} (e : Expr t) : Expr t
- := @InlineConstAndOp base_type op interp_base_type interp_op make_const t e.
-End Rewrite.
diff --git a/src/Compilers/Z/InlineConstAndOpByRewriteInterp.v b/src/Compilers/Z/InlineConstAndOpByRewriteInterp.v
deleted file mode 100644
index e67442c71..000000000
--- a/src/Compilers/Z/InlineConstAndOpByRewriteInterp.v
+++ /dev/null
@@ -1,12 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InlineConstAndOpByRewriteInterp.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.InlineConstAndOpByRewrite.
-
-Module Export Rewrite.
- Definition InterpInlineConstAndOp {t} (e : Expr t)
- : forall x, Interp (InlineConstAndOp e) x = Interp e x
- := @InterpInlineConstAndOp _ _ _ _ _ t e Syntax.Util.make_const_correct.
-
- Hint Rewrite @InterpInlineConstAndOp : reflective_interp.
-End Rewrite.
diff --git a/src/Compilers/Z/InlineConstAndOpByRewriteWf.v b/src/Compilers/Z/InlineConstAndOpByRewriteWf.v
deleted file mode 100644
index aa883f5e7..000000000
--- a/src/Compilers/Z/InlineConstAndOpByRewriteWf.v
+++ /dev/null
@@ -1,13 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.InlineConstAndOpByRewriteWf.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.InlineConstAndOpByRewrite.
-
-Module Export Rewrite.
- Definition Wf_InlineConstAndOp {t} (e : Expr t) (Hwf : Wf e)
- : Wf (InlineConstAndOp e)
- := @Wf_InlineConstAndOp _ _ _ _ _ t e Hwf.
-
- Hint Resolve Wf_InlineConstAndOp : wf.
-End Rewrite.
diff --git a/src/Compilers/Z/InlineConstAndOpInterp.v b/src/Compilers/Z/InlineConstAndOpInterp.v
deleted file mode 100644
index 417085929..000000000
--- a/src/Compilers/Z/InlineConstAndOpInterp.v
+++ /dev/null
@@ -1,11 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.InlineConstAndOpInterp.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.InlineConstAndOp.
-
-Definition InterpInlineConstAndOp {t} (e : Expr t) (Hwf : Wf e)
- : forall x, Interp (InlineConstAndOp e) x = Interp e x
- := @InterpInlineConstAndOp _ _ _ _ _ t e Hwf Syntax.Util.make_const_correct.
-
-Hint Rewrite @InterpInlineConstAndOp using solve_wf_side_condition : reflective_interp.
diff --git a/src/Compilers/Z/InlineConstAndOpWf.v b/src/Compilers/Z/InlineConstAndOpWf.v
deleted file mode 100644
index aa0f41d07..000000000
--- a/src/Compilers/Z/InlineConstAndOpWf.v
+++ /dev/null
@@ -1,11 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.InlineConstAndOpWf.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.InlineConstAndOp.
-
-Definition Wf_InlineConstAndOp {t} (e : Expr t) (Hwf : Wf e)
- : Wf (InlineConstAndOp e)
- := @Wf_InlineConstAndOp _ _ _ _ _ t e Hwf.
-
-Hint Resolve Wf_InlineConstAndOp : wf.
diff --git a/src/Compilers/Z/InlineInterp.v b/src/Compilers/Z/InlineInterp.v
deleted file mode 100644
index 6413a68cd..000000000
--- a/src/Compilers/Z/InlineInterp.v
+++ /dev/null
@@ -1,15 +0,0 @@
-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 InterpInlineConstAndOpp {interp_base_type interp_op} {t} (e : Expr t) (Hwf : Wf e)
- : forall x, Compilers.Syntax.Interp interp_op (InlineConstAndOpp e) x = Compilers.Syntax.Interp interp_op e x
- := @InterpInlineConst _ interp_base_type _ _ _ t e Hwf.
-
-Definition InterpInlineConst {interp_base_type interp_op} {t} (e : Expr t) (Hwf : Wf e)
- : forall x, Compilers.Syntax.Interp interp_op (InlineConst e) x = Compilers.Syntax.Interp interp_op e x
- := @InterpInlineConst _ interp_base_type _ _ _ t e Hwf.
-
-Hint Rewrite @InterpInlineConstAndOpp @InterpInlineConst using solve_wf_side_condition : reflective_interp.
diff --git a/src/Compilers/Z/InlineWf.v b/src/Compilers/Z/InlineWf.v
deleted file mode 100644
index b7726987b..000000000
--- a/src/Compilers/Z/InlineWf.v
+++ /dev/null
@@ -1,15 +0,0 @@
-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_InlineConstAndOpp {t} (e : Expr t) (Hwf : Wf e)
- : Wf (InlineConstAndOpp e)
- := @Wf_InlineConst _ _ _ t e Hwf.
-
-Definition Wf_InlineConst {t} (e : Expr t) (Hwf : Wf e)
- : Wf (InlineConst e)
- := @Wf_InlineConst _ _ _ t e Hwf.
-
-Hint Resolve Wf_InlineConstAndOpp Wf_InlineConst : wf.
diff --git a/src/Compilers/Z/InterpSideConditions.v b/src/Compilers/Z/InterpSideConditions.v
deleted file mode 100644
index bc3c8f6a9..000000000
--- a/src/Compilers/Z/InterpSideConditions.v
+++ /dev/null
@@ -1,9 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Util.
-Require Import Crypto.Compilers.InterpSideConditions.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Notations.
-
-Definition InterpSideConditions {t} (e : Expr t) : interp_flat_type interp_base_type (domain t) -> pointed_Prop
- := InterpSideConditions interp_op (@interped_op_side_conditions) e.
diff --git a/src/Compilers/Z/JavaNotations.v b/src/Compilers/Z/JavaNotations.v
deleted file mode 100644
index d7ce672c1..000000000
--- a/src/Compilers/Z/JavaNotations.v
+++ /dev/null
@@ -1,882 +0,0 @@
-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 "T0 x , T1 y = A ; b" (at level 200, b at level 200, format "T0 x , T1 y = A ; '//' b").
-Reserved Notation "v == 0 ? a : b" (at level 40, a at level 10, b at level 10).
-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.
-Notation "T0 x , T1 y = A ; b" := (LetIn (tx:=Prod T0 T1) A (fun '((x, y)%core) => 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:
-<<
-#!/usr/bin/env python
-# -*- coding: utf-8 -*-
-types = ('int', 'int', 'int', 'int', 'int', 'int', 'long', 'uint128_t', 'uint256_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, x at level 9, y at level 9).'
- % (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, y at level 9).'
- % ('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, x at level 9).'
- % (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 "'uint256_t'" := (Tbase (TWord 8)).
-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, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x * '(int)' y" := (Op (Mul (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(int)' y" := (Op (Mul (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x * y" := (Op (Mul (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(long)' x * '(long)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(long)' x * '(long)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(long)' x * '(long)' y" := (Op (Mul (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(long)' y" := (Op (Mul (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(long)' x * y" := (Op (Mul (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(uint128_t)' x * '(uint128_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(uint128_t)' y" := (Op (Mul (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(uint128_t)' x * y" := (Op (Mul (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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 "'(uint256_t)' x * '(uint256_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 40, x at level 9, y at level 9).
-Notation "'(uint256_t)' x * '(uint256_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(uint256_t)' x * '(uint256_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(uint256_t)' x * '(uint256_t)' y" := (Op (Mul (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x * '(uint256_t)' y" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(uint256_t)' x * y" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (at level 40, x at level 9).
-Notation "x * '(uint256_t)' y" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(uint256_t)' x * y" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x * '(uint256_t)' y" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(uint256_t)' x * y" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x * '(uint256_t)' y" := (Op (Mul (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(uint256_t)' x * y" := (Op (Mul (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-Notation "x * y" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (Pair x y)).
-Notation "x * y" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))).
-Notation "x * y" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)).
-Notation "x * y" := (Op (Mul (TWord 8) (TWord 8) (TWord 8)) (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, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x + '(int)' y" := (Op (Add (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(int)' y" := (Op (Add (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x + y" := (Op (Add (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(long)' x + '(long)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(long)' x + '(long)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(long)' x + '(long)' y" := (Op (Add (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(long)' y" := (Op (Add (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(long)' x + y" := (Op (Add (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(uint128_t)' x + '(uint128_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(uint128_t)' y" := (Op (Add (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(uint128_t)' x + y" := (Op (Add (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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 "'(uint256_t)' x + '(uint256_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 50, x at level 9, y at level 9).
-Notation "'(uint256_t)' x + '(uint256_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(uint256_t)' x + '(uint256_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(uint256_t)' x + '(uint256_t)' y" := (Op (Add (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x + '(uint256_t)' y" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(uint256_t)' x + y" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (at level 50, x at level 9).
-Notation "x + '(uint256_t)' y" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(uint256_t)' x + y" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x + '(uint256_t)' y" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(uint256_t)' x + y" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x + '(uint256_t)' y" := (Op (Add (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(uint256_t)' x + y" := (Op (Add (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-Notation "x + y" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (Pair x y)).
-Notation "x + y" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))).
-Notation "x + y" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)).
-Notation "x + y" := (Op (Add (TWord 8) (TWord 8) (TWord 8)) (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, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(int)' x - '(int)' y" := (Op (Sub (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(int)' y" := (Op (Sub (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(int)' x - y" := (Op (Sub (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(long)' x - '(long)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(long)' x - '(long)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(long)' x - '(long)' y" := (Op (Sub (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(long)' y" := (Op (Sub (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(long)' x - y" := (Op (Sub (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(uint128_t)' x - '(uint128_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(uint128_t)' y" := (Op (Sub (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(uint128_t)' x - y" := (Op (Sub (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-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 "'(uint256_t)' x - '(uint256_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 50, x at level 9, y at level 9).
-Notation "'(uint256_t)' x - '(uint256_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "'(uint256_t)' x - '(uint256_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9, y at level 9).
-Notation "'(uint256_t)' x - '(uint256_t)' y" := (Op (Sub (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9, y at level 9).
-Notation "x - '(uint256_t)' y" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (at level 50, y at level 9).
-Notation "'(uint256_t)' x - y" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (at level 50, x at level 9).
-Notation "x - '(uint256_t)' y" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 50, y at level 9).
-Notation "'(uint256_t)' x - y" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 50, x at level 9).
-Notation "x - '(uint256_t)' y" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 50, y at level 9).
-Notation "'(uint256_t)' x - y" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 50, x at level 9).
-Notation "x - '(uint256_t)' y" := (Op (Sub (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, y at level 9).
-Notation "'(uint256_t)' x - y" := (Op (Sub (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 50, x at level 9).
-Notation "x - y" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (Pair x y)).
-Notation "x - y" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))).
-Notation "x - y" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)).
-Notation "x - y" := (Op (Sub (TWord 8) (TWord 8) (TWord 8)) (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, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(int)' x & '(int)' y" := (Op (Land (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(int)' y" := (Op (Land (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(int)' x & y" := (Op (Land (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(long)' x & '(long)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(long)' x & '(long)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(long)' x & '(long)' y" := (Op (Land (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(long)' y" := (Op (Land (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(long)' x & y" := (Op (Land (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(uint128_t)' x & '(uint128_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(uint128_t)' y" := (Op (Land (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(uint128_t)' x & y" := (Op (Land (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-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 "'(uint256_t)' x & '(uint256_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 40, x at level 9, y at level 9).
-Notation "'(uint256_t)' x & '(uint256_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "'(uint256_t)' x & '(uint256_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9, y at level 9).
-Notation "'(uint256_t)' x & '(uint256_t)' y" := (Op (Land (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9, y at level 9).
-Notation "x & '(uint256_t)' y" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (at level 40, y at level 9).
-Notation "'(uint256_t)' x & y" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (at level 40, x at level 9).
-Notation "x & '(uint256_t)' y" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 40, y at level 9).
-Notation "'(uint256_t)' x & y" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 40, x at level 9).
-Notation "x & '(uint256_t)' y" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 40, y at level 9).
-Notation "'(uint256_t)' x & y" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 40, x at level 9).
-Notation "x & '(uint256_t)' y" := (Op (Land (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, y at level 9).
-Notation "'(uint256_t)' x & y" := (Op (Land (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 40, x at level 9).
-Notation "x & y" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (Pair x y)).
-Notation "x & y" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))).
-Notation "x & y" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)).
-Notation "x & y" := (Op (Land (TWord 8) (TWord 8) (TWord 8)) (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, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 0) (TWord _) (TWord 0)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 0) (TWord 0)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 1) (TWord _) (TWord 1)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 1) (TWord 1)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 2) (TWord _) (TWord 2)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 2) (TWord 2)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 3) (TWord _) (TWord 3)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 3) (TWord 3)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 4) (TWord _) (TWord 4)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 4) (TWord 4)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(int)' x << '(int)' y" := (Op (Shl (TWord _) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(int)' y" := (Op (Shl (TWord 5) (TWord _) (TWord 5)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(int)' x << y" := (Op (Shl (TWord _) (TWord 5) (TWord 5)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(long)' x << '(long)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(long)' x << '(long)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(long)' x << '(long)' y" := (Op (Shl (TWord _) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(long)' y" := (Op (Shl (TWord 6) (TWord _) (TWord 6)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(long)' x << y" := (Op (Shl (TWord _) (TWord 6) (TWord 6)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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, x at level 9, y at level 9).
-Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(uint128_t)' x << '(uint128_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(uint128_t)' y" := (Op (Shl (TWord 7) (TWord _) (TWord 7)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(uint128_t)' x << y" := (Op (Shl (TWord _) (TWord 7) (TWord 7)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-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 "'(uint256_t)' x << '(uint256_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 30, x at level 9, y at level 9).
-Notation "'(uint256_t)' x << '(uint256_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "'(uint256_t)' x << '(uint256_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 30, x at level 9, y at level 9).
-Notation "'(uint256_t)' x << '(uint256_t)' y" := (Op (Shl (TWord _) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 30, x at level 9, y at level 9).
-Notation "x << '(uint256_t)' y" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair x y)) (at level 30, y at level 9).
-Notation "'(uint256_t)' x << y" := (Op (Shl (TWord _) (TWord 8) (TWord 8)) (Pair x y)) (at level 30, x at level 9).
-Notation "x << '(uint256_t)' y" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 30, y at level 9).
-Notation "'(uint256_t)' x << y" := (Op (Shl (TWord _) (TWord 8) (TWord 8)) (Pair x (Var y))) (at level 30, x at level 9).
-Notation "x << '(uint256_t)' y" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 30, y at level 9).
-Notation "'(uint256_t)' x << y" := (Op (Shl (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) y)) (at level 30, x at level 9).
-Notation "x << '(uint256_t)' y" := (Op (Shl (TWord 8) (TWord _) (TWord 8)) (Pair (Var x) (Var y))) (at level 30, y at level 9).
-Notation "'(uint256_t)' x << y" := (Op (Shl (TWord _) (TWord 8) (TWord 8)) (Pair (Var x) (Var y))) (at level 30, x at level 9).
-Notation "x << y" := (Op (Shl (TWord 8) (TWord 8) (TWord 8)) (Pair x y)).
-Notation "x << y" := (Op (Shl (TWord 8) (TWord 8) (TWord 8)) (Pair x (Var y))).
-Notation "x << y" := (Op (Shl (TWord 8) (TWord 8) (TWord 8)) (Pair (Var x) y)).
-Notation "x << y" := (Op (Shl (TWord 8) (TWord 8) (TWord 8)) (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 "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (Pair x y)) (at level 30).
-Notation "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (Pair x (Var y))) (at level 30).
-Notation "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (Pair (Var x) y)) (at level 30).
-Notation "'(uint256_t)' ( x >> y )" := (Op (Shr (TWord _) (TWord _) (TWord 8)) (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
deleted file mode 100644
index 1985653d4..000000000
--- a/src/Compilers/Z/MapCastByDeBruijn.v
+++ /dev/null
@@ -1,28 +0,0 @@
-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
deleted file mode 100644
index 0103342f4..000000000
--- a/src/Compilers/Z/MapCastByDeBruijnInterp.v
+++ /dev/null
@@ -1,56 +0,0 @@
-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.Syntax.Util.
-Require Import Crypto.Compilers.Z.MapCastByDeBruijn.
-Require Import Crypto.Compilers.Z.InterpSideConditions.
-Require Import Crypto.Util.PointedProp.
-
-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)
- (Hside : to_prop (interped_op_side_conditions opc 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))
- (Hside : to_prop (interped_op_side_conditions opc (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 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)
- (Hside : to_prop (InterpSideConditions e v)),
- Compilers.Syntax.Interp interp_op_bounds e input_bounds = b
- /\ @inbounds _ b (Compilers.Syntax.Interp interp_op e v)
- /\ cast_back _ _ (Compilers.Syntax.Interp interp_op e' v') = (Compilers.Syntax.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
deleted file mode 100644
index 791e33886..000000000
--- a/src/Compilers/Z/MapCastByDeBruijnWf.v
+++ /dev/null
@@ -1,56 +0,0 @@
-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 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 (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/Named/DeadCodeElimination.v b/src/Compilers/Z/Named/DeadCodeElimination.v
deleted file mode 100644
index 449494e74..000000000
--- a/src/Compilers/Z/Named/DeadCodeElimination.v
+++ /dev/null
@@ -1,14 +0,0 @@
-Require Import Coq.PArith.BinPos.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.DeadCodeElimination.
-Require Import Crypto.Compilers.Z.Syntax.
-
-Section language.
- Context {Name : Type}
- {Context : Context Name (fun _ : base_type => positive)}.
-
- Definition EliminateDeadCode {t} e ls
- := @EliminateDeadCode base_type op Name _ internal_base_type_dec_bl Context t e ls.
- Definition CompileAndEliminateDeadCode {t} e ls
- := @CompileAndEliminateDeadCode base_type op Name _ internal_base_type_dec_bl Context t e ls.
-End language.
diff --git a/src/Compilers/Z/Named/DeadCodeEliminationInterp.v b/src/Compilers/Z/Named/DeadCodeEliminationInterp.v
deleted file mode 100644
index 422e8c309..000000000
--- a/src/Compilers/Z/Named/DeadCodeEliminationInterp.v
+++ /dev/null
@@ -1,50 +0,0 @@
-Require Import Coq.PArith.BinPos.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Z.Named.DeadCodeElimination.
-Require Import Crypto.Compilers.Named.DeadCodeEliminationInterp.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Named.Syntax.
-
-Section language.
- Context {interp_base_type : base_type -> Type}
- {interp_op : forall s d, op s d -> interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d}
- {Name : Type}
- {InContext : Context Name (fun _ : base_type => positive)}
- {InContextOk : ContextOk InContext}
- {Name_beq : Name -> Name -> bool}
- {InterpContext : Context Name interp_base_type}
- {InterpContextOk : ContextOk InterpContext}
- (Name_bl : forall n1 n2, Name_beq n1 n2 = true -> n1 = n2)
- (Name_lb : forall n1 n2, n1 = n2 -> Name_beq n1 n2 = true).
-
- Local Notation EliminateDeadCode := (@EliminateDeadCode Name InContext).
- Local Notation PContext var := (@PositiveContext base_type var base_type_beq internal_base_type_dec_bl).
-
- Lemma interp_EliminateDeadCode
- t e new_names
- (ctxi_interp : PContext _)
- (ctxr_interp : InterpContext)
- eout v1 v2 x
- : @EliminateDeadCode t e new_names = Some eout
- -> interp (interp_op:=interp_op) (ctx:=ctxr_interp) eout x = Some v1
- -> interp (interp_op:=interp_op) (ctx:=ctxi_interp) e x = Some v2
- -> v1 = v2.
- Proof using InContextOk InterpContextOk Name_bl Name_lb.
- eapply interp_EliminateDeadCode; eauto using internal_base_type_dec_lb.
- Qed.
-
- Lemma InterpEliminateDeadCode
- t e new_names
- eout
- v1 v2 x
- : @EliminateDeadCode t e new_names = Some eout
- -> Interp (Context:=InterpContext) (interp_op:=interp_op) eout x = Some v1
- -> Interp (Context:=PContext _) (interp_op:=interp_op) e x = Some v2
- -> v1 = v2.
- Proof using InContextOk InterpContextOk Name_bl Name_lb.
- apply interp_EliminateDeadCode.
- Qed.
-End language.
diff --git a/src/Compilers/Z/Named/RewriteAddToAdc.v b/src/Compilers/Z/Named/RewriteAddToAdc.v
deleted file mode 100644
index 784f7003c..000000000
--- a/src/Compilers/Z/Named/RewriteAddToAdc.v
+++ /dev/null
@@ -1,147 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Named.GetNames.
-Require Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Notations.
-
-Local Open Scope Z_scope.
-
-Section named.
- Context {Name : Type}
- (name_beq : Name -> Name -> bool).
- Import Named.Syntax.
- Local Notation flat_type := (flat_type base_type).
- Local Notation type := (type base_type).
- Local Notation exprf := (@exprf base_type op Name).
- Local Notation expr := (@expr base_type op Name).
-
- Local Notation tZ := (Tbase TZ).
- Local Notation ADC bw c x y := (Op (@AddWithGetCarry bw TZ TZ TZ TZ TZ)
- (Pair (Pair (t1:=tZ) c (t2:=tZ) x) (t2:=tZ) y)).
- Local Notation ADD bw x y := (ADC bw (Op (OpConst 0) TT) x y).
- Local Notation ADX x y := (Op (@Add TZ TZ TZ) (Pair (t1:=tZ) x (t2:=tZ) y)).
- Local Infix "=Z?" := Z.eqb.
- Local Infix "=n?" := name_beq.
-
- Definition is_const_or_var {t} (v : exprf t)
- := match v with
- | Var _ _ => true
- | Op _ _ (OpConst _ _) TT => true
- | _ => false
- end.
-
- Fixpoint name_list_has_duplicate (ls : list Name) : bool
- := match ls with
- | nil => false
- | cons n ns
- => orb (name_list_has_duplicate ns)
- (List.fold_left orb (List.map (name_beq n) ns) false)
- end.
-
- Definition invertT t
- := option ((Name * Name * Z * exprf tZ * exprf tZ)
- * (Name * Name * Z * exprf tZ * Name)
- * (((Name * Name * Name) * exprf t)
- + exprf t)).
-
- Definition invert_for_do_rewrite_step1 {t} (e : exprf t)
- : option ((Name * Name * Z * exprf tZ * exprf tZ) * exprf t)
- := match e in Named.exprf _ _ _ t return option ((Name * Name * Z * exprf tZ * exprf tZ) * exprf t) with
- | (nlet (a2, c1) : tZ * tZ := (ADD bw1 a b as ex0) in P0)
- => Some ((a2, c1, bw1, a, b), P0)
- | _ => None
- end%core%nexpr%bool.
- Definition invert_for_do_rewrite_step2 {t} (e : exprf t)
- : option ((Name * Name * Z * exprf tZ * Name) * exprf t)
- := match e in Named.exprf _ _ _ t return option ((Name * Name * Z * exprf tZ * Name) * exprf t) with
- | (nlet (s , c2) : tZ * tZ := (ADD bw2 c0 (Var TZ a2') as ex1) in P1)
- => Some ((s, c2, bw2, c0, a2'), P1)
- | _ => None
- end%core%nexpr%bool.
- Definition invert_for_do_rewrite_step3 {t} (e : exprf t)
- : option ((Name * Name * Name) * exprf t)
- := match e in Named.exprf _ _ _ t return option ((Name * Name * Name) * exprf t) with
- | (nlet c : tZ := (ADX (Var TZ c1') (Var TZ c2') as ex2) in P)
- => Some ((c, c1', c2'), P)
- | _ => None
- end%core%nexpr%bool.
-
- Definition invert_for_do_rewrite {t} (e : exprf t)
- : invertT t
- := match invert_for_do_rewrite_step1 e with
- | Some ((a2, c1, bw1, a, b), P0) (* (nlet (a2, c1) : tZ * tZ := (ADD bw1 a b as ex0) in P0) *)
- => match invert_for_do_rewrite_step2 P0 with
- | Some ((s, c2, bw2, c0, a2'), P1) (* (nlet (s , c2) : tZ * tZ := (ADD bw2 c0 (Var TZ a2') as ex1) in P1) *)
- => match match invert_for_do_rewrite_step3 P1 with
- | Some ((c, c1', c2'), P) (* (nlet c : tZ := (ADX (Var TZ c1') (Var TZ c2') as ex2) in P) as P1' *)
- => if (((bw1 =Z? bw2) && (a2 =n? a2') && (c1 =n? c1') && (c2 =n? c2'))
- && (is_const_or_var c0 && is_const_or_var a && is_const_or_var b)
- && negb (name_list_has_duplicate (a2::c1::s::c2::c::nil ++ get_namesf c0 ++ get_namesf a ++ get_namesf b)%list))
- then Some ((a2, c1, bw1, a, b),
- (s, c2, bw2, c0, a2'),
- inl ((c, c1', c2'), P))
- else None
- | None => None
- end with
- | Some v => Some v
- | None => if (((bw1 =Z? bw2) && (a2 =n? a2'))
- && (is_const_or_var c0 && is_const_or_var a && is_const_or_var b)
- && negb (name_list_has_duplicate (a2::c1::s::c2::nil ++ get_namesf c0 ++ get_namesf a ++ get_namesf b)%list))
- then Some ((a2, c1, bw1, a, b),
- (s, c2, bw2, c0, a2'),
- inr P1)
- else None
- end
- | None => None
- end
- | None => None
- end%core%nexpr%bool.
-
- Definition do_rewrite {t} (e : exprf t)
- : exprf t
- := match invert_for_do_rewrite e with
- | Some ((a2, c1, bw1, a, b),
- (s, c2, bw2, c0, a2'),
- inl ((c, c1', c2'), P))
- => (nlet (a2, c1) : tZ * tZ := ADD bw1 a b in
- nlet (s , c2) : tZ * tZ := ADD bw2 c0 (Var a2') in
- nlet c : tZ := ADX (Var c1') (Var c2') in
- nlet (s, c) : tZ * tZ := ADC bw1 c0 a b in
- P)
- | Some ((a2, c1, bw1, a, b),
- (s, c2, bw2, c0, a2'),
- inr P)
- => (nlet (a2, c1) : tZ * tZ := ADD bw1 a b in
- nlet (s , c2) : tZ * tZ := ADD bw2 c0 (Var a2') in
- nlet s : tZ := (nlet (s, c1) : tZ * tZ := ADC bw1 c0 a b in Var s) in
- P)
- | None
- => e
- end%core%nexpr.
-
- Definition rewrite_exprf_prestep
- (rewrite_exprf : forall {t} (e : exprf t), exprf t)
- {t} (e : exprf t)
- : exprf t
- := match e in Named.exprf _ _ _ t return exprf t with
- | TT => TT
- | Var _ n => Var n
- | Op _ _ opc args
- => Op opc (@rewrite_exprf _ args)
- | (nlet nx := ex in eC)
- => (nlet nx := @rewrite_exprf _ ex in @rewrite_exprf _ eC)
- | Pair tx ex ty ey
- => Pair (@rewrite_exprf tx ex) (@rewrite_exprf ty ey)
- end%nexpr.
-
- Fixpoint rewrite_exprf {t} (e : exprf t) : exprf t
- := do_rewrite (@rewrite_exprf_prestep (@rewrite_exprf) t e).
-
- Definition rewrite_expr {t} (e : expr t) : expr t
- := match e in Named.expr _ _ _ t return expr t with
- | Abs _ _ n f => Abs n (rewrite_exprf f)
- end.
-End named.
diff --git a/src/Compilers/Z/Named/RewriteAddToAdcInterp.v b/src/Compilers/Z/Named/RewriteAddToAdcInterp.v
deleted file mode 100644
index 8f4dc4644..000000000
--- a/src/Compilers/Z/Named/RewriteAddToAdcInterp.v
+++ /dev/null
@@ -1,449 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.ContextProperties.Proper.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Equality.
-Require Import Crypto.Compilers.Z.Named.RewriteAddToAdc.
-Require Import Crypto.Compilers.Named.GetNames.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Sum.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.ZUtil.AddGetCarry.
-Require Import Crypto.Util.ListUtil.FoldBool.
-Require Import Crypto.Util.Decidable.
-
-Local Open Scope Z_scope.
-
-Section named.
- Context {Name : Type}
- {InterpContext : Context Name interp_base_type}
- {InterpContextOk : ContextOk InterpContext}
- (Name_beq : Name -> Name -> bool)
- (Name_bl : forall n1 n2, Name_beq n1 n2 = true -> n1 = n2)
- (Name_lb : forall n1 n2, n1 = n2 -> Name_beq n1 n2 = true).
-
- Local Notation name_list_has_duplicate := (@name_list_has_duplicate Name Name_beq).
- Local Notation exprf := (@exprf base_type op Name).
- Local Notation expr := (@expr base_type op Name).
- Local Notation do_rewrite := (@do_rewrite Name Name_beq).
- Local Notation invert_for_do_rewrite := (@invert_for_do_rewrite Name Name_beq).
- Local Notation rewrite_exprf := (@rewrite_exprf Name Name_beq).
- Local Notation rewrite_exprf_prestep := (@rewrite_exprf_prestep Name).
- Local Notation rewrite_expr := (@rewrite_expr Name Name_beq).
-
- Local Instance Name_dec : DecidableRel (@eq Name)
- := dec_rel_of_bool_dec_rel Name_beq Name_bl Name_lb.
-
- Local Notation retT e re :=
- (forall (ctx : InterpContext)
- v,
- Named.interpf (interp_op:=interp_op) (ctx:=ctx) re = Some v
- -> Named.interpf (interp_op:=interp_op) (ctx:=ctx) e = Some v)
- (only parsing).
- Local Notation tZ := (Tbase TZ).
- Local Notation ADC bw c x y := (Op (@AddWithGetCarry bw TZ TZ TZ TZ TZ)
- (Pair (Pair (t1:=tZ) c (t2:=tZ) x) (t2:=tZ) y)).
- Local Notation ADD bw x y := (ADC bw (Op (OpConst 0) TT) x y).
- Local Notation ADX x y := (Op (@Add TZ TZ TZ) (Pair (t1:=tZ) x (t2:=tZ) y)).
- Local Infix "=Z?" := Z.eqb.
- Local Infix "=n?" := Name_beq.
-
- Local Ltac simple_t_step :=
- first [ exact I
- | progress intros
- | progress subst
- | progress inversion_option
- | progress inversion_sum
- | progress inversion_prod ].
- Local Ltac destruct_t_step :=
- first [ break_innermost_match_hyps_step
- | break_innermost_match_step ].
- Local Ltac do_small_inversion e :=
- is_var e;
- lazymatch type of e with
- | exprf ?T
- => revert dependent e;
- let P := match goal with |- forall e, @?P e => P end in
- intro e;
- lazymatch T with
- | Unit
- => refine match e in Named.exprf _ _ _ t return match t return Named.exprf _ _ _ t -> _ with Unit => P | _ => fun _ => True end e with TT => _ | _ => _ end
- | tZ
- => refine match e in Named.exprf _ _ _ t return match t return Named.exprf _ _ _ t -> _ with tZ => P | _ => fun _ => True end e with TT => _ | _ => _ end
- | (tZ * tZ)%ctype
- => refine match e in Named.exprf _ _ _ t return match t return Named.exprf _ _ _ t -> _ with (tZ * tZ)%ctype => P | _ => fun _ => True end e with TT => _ | _ => _ end
- | (tZ * tZ * tZ)%ctype
- => refine match e in Named.exprf _ _ _ t return match t return Named.exprf _ _ _ t -> _ with (tZ * tZ * tZ)%ctype => P | _ => fun _ => True end e with TT => _ | _ => _ end
- end;
- try exact I
- | op ?a ?T
- => first [ is_var a;
- move e at top;
- revert dependent a;
- let P := match goal with |- forall a e, @?P a e => P end in
- intros a e;
- lazymatch T with
- | tZ
- => refine match e in op a t return match t return op a t -> _ with tZ => P a | _ => fun _ => True end e with OpConst _ _ => _ | _ => _ end
- | (tZ * tZ)%ctype
- => refine match e in op a t return match t return op a t -> _ with (tZ * tZ)%ctype => P a | _ => fun _ => True end e with OpConst _ _ => _ | _ => _ end
- end ];
- try exact I
- end.
- Local Ltac small_inversion_prestep _ :=
- match goal with
- | [ H : match ?e with _ => _ end = Some _ |- _ ] => do_small_inversion e
- | [ H : match ?e with _ => _ end = true |- _ ] => do_small_inversion e
- | [ H : match ?e with _ => _ end _ = Some _ |- _ ] => do_small_inversion e
- end.
- Local Ltac small_inversion_step :=
- small_inversion_prestep (); break_innermost_match; intros; try exact I.
-
- Local Ltac t_rewrite_step_correct_step :=
- first [ reflexivity
- | progress inversion_option
- | simple_t_step
- | break_innermost_match_hyps_step
- | small_inversion_step
- | progress unfold invert_for_do_rewrite_step1, invert_for_do_rewrite_step2, invert_for_do_rewrite_step3 in * ].
- Local Ltac t_rewrite_step_correct := repeat t_rewrite_step_correct_step.
-
- Local Ltac mk_lookupb_extendb_lemma_debug := constr:(false).
- Local Ltac debug_print_fail tac :=
- let lvl := mk_lookupb_extendb_lemma_debug in
- lazymatch lvl with
- | true => let dummy := match goal with
- | _ => tac ()
- end in
- constr:(I : I)
- | false => constr:(I : I)
- | _ => let TRUE := uconstr:(true) in
- let FALSE := uconstr:(false) in
- let dummy := match goal with
- | _ => idtac "Error: Invalid mk_lookupb_extendb_lemma_debug level" lvl "which is neither" TRUE "nor" FALSE
- end in
- constr:(I : I)
- end.
-
- (** We build the lemma explicitly, because letting [rewrite] and
- [assumption || congruence] pick out the hypotheses and build the
- lemmas is actually a bottleneck (timesavings: about 25s) *)
- Local Ltac mk_lookupb_extendb_lemma base_type_code Name var Context t t' ctx n n' v :=
- lazymatch n with
- | n' => lazymatch t with
- | t' => constr:(@lookupb_extendb_same base_type_code Name var Context _ ctx n t v)
- | _ => let lem := constr:(@lookupb_extendb_wrong_type base_type_code Name var Context _ ctx n t t' v) in
- lazymatch goal with
- | [ H : t <> t' |- _ ]
- => constr:(lem H)
- | [ H : t' <> t |- _ ]
- => constr:(lem (@not_eq_sym _ _ _ H))
- | _ => let HT := uconstr:(t <> t') in
- debug_print_fail ltac:(fun _ => idtac "Error in mk_lookupb_exntedb_lemma: could not find hypothesis of type" HT)
- end
- end
- | _ => let lem := constr:(@lookupb_extendb_different base_type_code Name var Context _ ctx n n' t t' v) in
- lazymatch goal with
- | [ H : n <> n' |- _ ]
- => constr:(lem H)
- | [ H : n' <> n |- _ ]
- => constr:(lem (@not_eq_sym _ _ _ H))
- | _ => let HT := uconstr:(n <> n') in
- debug_print_fail ltac:(fun _ => idtac "Error in mk_lookupb_exntedb_lemma: could not find hypothesis of type" HT)
- end
- end.
- Local Ltac rewrite_lookupb_step :=
- first [ match goal with
- | [ H : context[@lookupb ?base_type_code ?Name ?var ?Context ?t' (extendb (t:=?t) ?ctx ?n ?v) ?n'] |- _ ]
- => let lem := mk_lookupb_extendb_lemma base_type_code Name var Context t t' ctx n n' v in
- rewrite lem in H
- | [ |- context[@lookupb ?base_type_code ?Name ?var ?Context ?t' (extendb (t:=?t) ?ctx ?n ?v) ?n'] ]
- => let lem := mk_lookupb_extendb_lemma base_type_code Name var Context t t' ctx n n' v in
- rewrite lem
- end
- | match goal with
- | [ H : context[lookupb (extendb _ _ _) _] |- _ ] => revert H
- | [ |- context[lookupb (extendb _ ?n _) ?n'] ]
- => (tryif constr_eq n n' then fail else idtac);
- lazymatch goal with
- | [ H : n = n' |- _ ] => fail
- | [ H : n' = n |- _ ] => fail
- | [ H : n <> n' |- _ ] => fail
- | [ H : n' <> n |- _ ] => fail
- | _ => destruct (dec (n = n')); subst
- end
- | [ |- context[lookupb (t:=?t0) (extendb (t:=?t1) _ _ _) _] ]
- => (tryif constr_eq t0 t1 then fail else idtac);
- lazymatch goal with
- | [ H : t0 = t1 |- _ ] => fail
- | [ H : t1 = t0 |- _ ] => fail
- | [ H : t0 <> t1 |- _ ] => fail
- | [ H : t1 <> t0 |- _ ] => fail
- | _ => destruct (dec (t0 = t1)); subst
- end
- end ].
- Local Ltac rewrite_lookupb := repeat rewrite_lookupb_step.
-
- Local Ltac do_rewrite_adc' P :=
- let lem := open_constr:(Z.add_get_carry_to_add_with_get_carry_cps _ _ _ _ P) in
- let T := type of lem in
- let T := (eval cbv [Let_In Definitions.Z.add_with_get_carry Definitions.Z.add_with_get_carry Definitions.Z.get_carry Definitions.Z.add_get_carry] in T) in
- etransitivity; [ | eapply (lem : T) ];
- try reflexivity.
- Local Ltac do_rewrite_adc :=
- first [ do_rewrite_adc' uconstr:(fun a b => Some b)
- | do_rewrite_adc' uconstr:(fun a b => Some a) ].
- Local Ltac do_small_inversion_ctx :=
- repeat match goal with
- | [ H : is_const_or_var ?e = true |- _ ]
- => do_small_inversion e; break_innermost_match; intros; try exact I;
- simpl in H; try solve [ clear -H; discriminate ]
- | [ H : match ?e with _ => _ end = true |- _ ]
- => do_small_inversion e; break_innermost_match; intros; try exact I;
- simpl in H; try solve [ clear -H; discriminate ]
- | [ H : match ?e with _ => _ end _ = true |- _ ]
- => do_small_inversion e; break_innermost_match; intros; try exact I;
- simpl in H; try solve [ clear -H; discriminate ]
- end.
- Local Ltac t_fin_step :=
- match goal with
- | [ |- ?x = ?x ] => reflexivity
- | [ H : ?x = Some _ |- context[?x] ] => rewrite H
- | [ H : ?x = None |- context[?x] ] => rewrite H
- | [ H : ?x = Some ?a, H' : ?x = Some ?b |- _ ]
- => assert (a = b) by congruence; (subst a || subst b)
- | [ H : interpf ?x = Some ?v, H' : interpf ?x = None |- interpf _ = None ]
- => cut (Some v = None);
- [ congruence | rewrite <- H, <- H'; clear H H' ]
- | _ => progress rewrite_lookupb
- | _ => progress simpl in *
- | _ => progress intros
- | _ => progress subst
- | _ => progress inversion_option
- | [ |- (dlet x := _ in _) = (dlet y := _ in _) ]
- => apply Proper_Let_In_nd_changebody_eq; intros ??
- | _ => progress unfold Let_In
- | [ |- interpf ?x = interpf ?x ]
- => eapply @interpf_Proper; [ eauto with typeclass_instances.. | intros ?? | reflexivity ]
- | _ => progress break_innermost_match; try reflexivity
- | _ => progress break_innermost_match_hyps; try reflexivity
- | _ => progress break_match; try reflexivity
- end.
- Local Ltac t_fin :=
- repeat t_fin_step;
- try do_rewrite_adc.
-
-
- Lemma invert_for_do_rewrite_step1_correct {t} {e : exprf t} {v}
- (H : invert_for_do_rewrite_step1 e = Some v)
- : e = let '((a2, c1, bw1, a, b), P) := v in
- (nlet (a2, c1) : tZ * tZ := ADD bw1 a b in P)%nexpr.
- Proof. t_rewrite_step_correct. Qed.
- Lemma invert_for_do_rewrite_step2_correct {t} {e : exprf t} {v}
- (H : invert_for_do_rewrite_step2 e = Some v)
- : e = let '((s, c2, bw2, c0, a2'), P) := v in
- (nlet (s , c2) : tZ * tZ := (ADD bw2 c0 (Var a2')) in P)%nexpr.
- Proof. t_rewrite_step_correct. Qed.
- Lemma invert_for_do_rewrite_step3_correct {t} {e : exprf t} {v}
- (H : invert_for_do_rewrite_step3 e = Some v)
- : e = let '((c, c1', c2'), P) := v in
- (nlet c : tZ := (ADX (Var c1') (Var c2')) in P)%nexpr.
- Proof. t_rewrite_step_correct. Qed.
-
- Lemma invert_for_do_rewrite_correct {t} {e : exprf t} {v}
- (H : invert_for_do_rewrite e = Some v)
- : match v with
- | ((a2, c1, bw1, a, b),
- (s, c2, bw2, c0, a2'),
- inl ((c, c1', c2'), P))
- => (nlet (a2, c1) : tZ * tZ := ADD bw1 a b in
- nlet (s , c2) : tZ * tZ := ADD bw2 c0 (Var a2') in
- nlet c : tZ := ADX (Var c1') (Var c2') in
- P) = e
- /\ ((((bw1 =Z? bw2) && (a2 =n? a2') && (c1 =n? c1') && (c2 =n? c2'))
- && (is_const_or_var c0 && is_const_or_var a && is_const_or_var b)
- && negb (name_list_has_duplicate (a2::c1::s::c2::c::nil ++ get_namesf c0 ++ get_namesf a ++ get_namesf b)%list))
- = true)
- | ((a2, c1, bw1, a, b),
- (s, c2, bw2, c0, a2'),
- inr P)
- => (nlet (a2, c1) : tZ * tZ := ADD bw1 a b in
- nlet (s , c2) : tZ * tZ := ADD bw2 c0 (Var a2') in
- P) = e
- /\ ((((bw1 =Z? bw2) && (a2 =n? a2'))
- && (is_const_or_var c0 && is_const_or_var a && is_const_or_var b)
- && negb (name_list_has_duplicate (a2::c1::s::c2::nil ++ get_namesf c0 ++ get_namesf a ++ get_namesf b)%list))
- = true)
- end%core%nexpr%bool.
- Proof.
- unfold invert_for_do_rewrite in H; break_innermost_match_hyps;
- repeat first [ progress subst
- | progress inversion_option
- | progress inversion_sum
- | progress inversion_prod
- | match goal with
- | [ H : _ = Some _ |- _ ]
- => first [ rewrite (invert_for_do_rewrite_step1_correct H)
- | rewrite (invert_for_do_rewrite_step2_correct H)
- | rewrite (invert_for_do_rewrite_step3_correct H) ]
- | [ H : ?x = true |- _ ] => rewrite H; clear H
- | [ |- _ /\ _ ] => split
- | [ |- ?x = ?x ] => reflexivity
- | [ H : Name_beq _ _ = true |- _ ] => apply Name_bl in H
- | [ H : Z.eqb _ _ = true |- _ ] => apply Z.eqb_eq in H
- | [ H : Name_beq ?x ?y = false |- _ ]
- => assert (x <> y) by (clear -H Name_lb; intro; rewrite Name_lb in H by assumption; congruence);
- clear H
- end
- | progress rewrite !Bool.negb_orb in *
- | progress rewrite !Bool.negb_true_iff in *
- | progress split_andb
- | progress simpl @negb in *
- | progress do_small_inversion_ctx ].
- Qed.
-
- Lemma interpf_do_rewrite
- {t} {e : exprf t}
- : retT e (do_rewrite e).
- Proof.
- unfold do_rewrite.
- pose proof (@invert_for_do_rewrite_correct t e) as H'.
- break_innermost_match; inversion_option;
- [ specialize (H' _ eq_refl)
- | specialize (H' _ eq_refl)
- | intros; subst; assumption ].
- all: intros *; let H := fresh in intro H; rewrite <- H; clear H.
- Time all:repeat first [ progress cbv beta iota in *
- | progress destruct_head'_and
- | progress subst
- | progress rewrite !Bool.negb_orb in *
- | progress split_andb
- | progress simpl @name_list_has_duplicate in *
- | match goal with
- | [ H : Name_beq _ _ = true |- _ ] => apply Name_bl in H
- | [ H : Z.eqb _ _ = true |- _ ] => apply Z.eqb_eq in H
- | [ H : Name_beq ?x ?y = false |- _ ]
- => assert (x <> y) by (clear -H Name_lb; intro; rewrite Name_lb in H by assumption; congruence);
- clear H
- | [ H : context[List.fold_left orb ?ls ?v] |- _ ]
- => lazymatch v with
- | false => fail
- | _ => rewrite (fold_left_orb_pull ls v) in H
- end
- end
- | progress simpl @List.fold_left in *
- | progress do_small_inversion_ctx
- | progress rewrite !Bool.negb_true_iff in *
- | progress intros *
- | match goal with
- | [ H : invert_for_do_rewrite _ = Some _ |- _ ] => clear H
- end ].
- Time all: repeat first [ progress simpl
- | progress cbv [interp_op option_map lift_op Zinterp_op] in * ].
- Time all: repeat first [ progress subst
- | progress inversion_option
- | progress rewrite_lookupb
- | progress intros
- | match goal with
- | [ H : ?x = Some ?a, H' : ?x = Some ?b |- _ ]
- => assert (a = b) by congruence; (subst a || subst b)
- | [ H : ?T, H' : ?T |- _ ] => clear H'
- end ].
- Set Ltac Profiling.
- Time all: repeat first [ reflexivity
- | progress subst
- | progress inversion_option
- | progress rewrite_lookupb
- | progress intros
- | progress cbn [fst snd]
- | match goal with
- | [ H : ?x = Some ?a, H' : ?x = Some ?b |- _ ]
- => assert (a = b) by congruence; (subst a || subst b)
- | [ H : ?T, H' : ?T |- _ ] => clear H'
- | [ |- context[match lookupb ?ctx ?n with _ => _ end] ]
- => is_var ctx; destruct (lookupb ctx n) eqn:?
- | [ |- (dlet x := ?e in _) = (dlet y := ?e in _) ]
- => apply Proper_Let_In_nd_changebody_eq; intros ??
- end
- | progress unfold Let_In at 1 ].
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- { Time t_fin. }
- Time Qed.
-
- Local Opaque RewriteAddToAdc.do_rewrite.
- Lemma interpf_rewrite_exprf
- {t} (e : exprf t)
- : retT e (rewrite_exprf e).
- Proof.
- pose t as T.
- pose (rewrite_exprf_prestep (@rewrite_exprf) e) as E.
- induction e; simpl in *;
- intros ctx v H;
- pose proof (interpf_do_rewrite (t:=T) (e:=E) ctx v H); clear H;
- subst T E;
- repeat first [ assumption
- | progress unfold option_map, Let_In in *
- | progress simpl in *
- | progress subst
- | progress inversion_option
- | apply (f_equal (@Some _))
- | break_innermost_match_step
- | break_innermost_match_hyps_step
- | congruence
- | solve [ eauto ]
- | match goal with
- | [ IH : forall ctx v, interpf ?e = Some v -> _ = Some _, H' : interpf ?e = Some _ |- _ ]
- => specialize (IH _ _ H')
- | [ H : ?x = Some ?a, H' : ?x = Some ?b |- _ ]
- => assert (a = b) by congruence; (subst a || subst b)
- | [ |- ?rhs = Some _ ]
- => lazymatch rhs with
- | Some _ => fail
- | None => fail
- | _ => destruct rhs eqn:?
- end
- end ].
- Qed.
-
- Lemma interp_rewrite_expr
- {t} (e : expr t)
- : forall (ctx : InterpContext)
- v x,
- Named.interp (interp_op:=interp_op) (ctx:=ctx) (rewrite_expr e) x = Some v
- -> Named.interp (interp_op:=interp_op) (ctx:=ctx) e x = Some v.
- Proof.
- unfold Named.interp, rewrite_expr; destruct e; simpl.
- intros *; apply interpf_rewrite_exprf.
- Qed.
-
- Lemma Interp_rewrite_expr
- {t} (e : expr t)
- : forall v x,
- Named.Interp (Context:=InterpContext) (interp_op:=interp_op) (rewrite_expr e) x = Some v
- -> Named.Interp (Context:=InterpContext) (interp_op:=interp_op) e x = Some v.
- Proof.
- intros *; apply interp_rewrite_expr.
- Qed.
-End named.
diff --git a/src/Compilers/Z/OpInversion.v b/src/Compilers/Z/OpInversion.v
deleted file mode 100644
index c27f51aee..000000000
--- a/src/Compilers/Z/OpInversion.v
+++ /dev/null
@@ -1,29 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.TypeInversion.
-Require Import Crypto.Compilers.Z.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
- | [ |- context[match ?e with OpConst _ _ => _ | _ => _ end] ]
- => invert_one_op e
- | [ H : context[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
deleted file mode 100644
index 7cb978241..000000000
--- a/src/Compilers/Z/Reify.v
+++ /dev/null
@@ -1,108 +0,0 @@
-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.Eta.
-Require Import Crypto.Compilers.EtaInterp.
-Require Import Crypto.Util.ZUtil.Definitions.
-Require Import Crypto.Util.IdfunWithAlt.
-Require Import Crypto.Util.SideConditions.CorePackages.
-Require Import Crypto.Util.Tactics.DebugPrint.
-
-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))
- | @Z.opp => constr:(reify_op op op_head 1 (Opp TZ TZ))
- | @Z.zselect => constr:(reify_op op op_head 3 (Zselect TZ TZ TZ TZ))
- | @Z.add_with_carry => constr:(reify_op op op_head 3 (AddWithCarry TZ TZ TZ TZ))
- | @Z.sub_with_borrow => constr:(reify_op op op_head 3 (SubWithBorrow TZ TZ TZ TZ))
- | @id_with_alt
- => lazymatch extra with
- | @id_with_alt Z _ _
- => constr:(reify_op op op_head 2 (IdWithAlt TZ TZ TZ))
- | @id_with_alt ?T _ _
- => let c := uconstr:(@id_with_alt) in
- let uZ := uconstr:(Z) in
- constr_run_tac_fail ltac:(fun _ => idtac "Error: In Reflection.Z.base_reify_op: can only reify" c "applied to" uZ "not" T)
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Anomaly: In Reflection.Z.base_reify_op: head is id_with_alt but body is wrong:" extra)
- end
- | @Z.mul_split_at_bitwidth
- => lazymatch extra with
- | @Z.mul_split_at_bitwidth ?bit_width _ _
- => constr:(reify_op op op_head 2 (MulSplit bit_width TZ TZ TZ TZ))
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Anomaly: In Reflection.Z.base_reify_op: head is Z.mul_split_with_bitwidth but body is wrong:" extra)
- end
- | @Z.add_with_get_carry
- => lazymatch extra with
- | @Z.add_with_get_carry ?bit_width _ _ _
- => constr:(reify_op op op_head 3 (AddWithGetCarry bit_width TZ TZ TZ TZ TZ))
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Anomaly: In Reflection.Z.base_reify_op: head is Z.add_with_get_carry but body is wrong:" extra)
- end
- | @Z.sub_with_get_borrow
- => lazymatch extra with
- | @Z.sub_with_get_borrow ?bit_width _ _ _
- => constr:(reify_op op op_head 3 (SubWithGetBorrow bit_width TZ TZ TZ TZ TZ))
- | _ => constr_run_tac_fail ltac:(fun _ => idtac "Anomaly: In Reflection.Z.base_reify_op: head is Z.sub_with_get_borrow but body is wrong:" extra)
- end
- end.
-Ltac base_reify_type T ::=
- lazymatch T with
- | Z => TZ
- | Z.Syntax.interp_base_type TZ => TZ
- | ?T
- => lazymatch (eval hnf in T) with
- | Z => TZ
- end
- end.
-Ltac Reify' e :=
- let e := (eval cbv beta delta [Z.add_get_carry Z.sub_get_borrow] in e) in
- Compilers.Reify.Reify' base_type interp_base_type op e.
-Ltac Reify e :=
- let e := (eval cbv beta delta [Z.add_get_carry Z.sub_get_borrow] in e) in
- 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 ()).
-
-Definition Reify_evar_package {t} f
- := @Reify_evar_package base_type interp_base_type op make_const interp_op t f.
-
-Definition Interp_Reify_evar_package
- {t f}
- (pkg : @Reify_evar_package t f)
- : forall x, Interp (val pkg) x = f x
- := evar_package_pf pkg.
-
-Ltac autosolve else_tac :=
- lazymatch goal with
- | [ |- @Reify_evar_package _ _ ]
- => eexists; cbv beta; Reify_rhs; reflexivity
- | _ => Compilers.Reify.autosolve else_tac
- end.
-
-Ltac SideConditions.CorePackages.autosolve ::= autosolve.
-
-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/RewriteAddToAdc.v b/src/Compilers/Z/RewriteAddToAdc.v
deleted file mode 100644
index 838c92b75..000000000
--- a/src/Compilers/Z/RewriteAddToAdc.v
+++ /dev/null
@@ -1,58 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Compilers.Named.Context.
-Require Import Crypto.Compilers.Named.Syntax.
-Require Import Crypto.Compilers.Named.InterpretToPHOAS.
-Require Import Crypto.Compilers.Named.Compile.
-Require Import Crypto.Compilers.Named.Wf.
-Require Import Crypto.Compilers.Named.CountLets.
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.Defaults.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Named.DeadCodeElimination.
-Require Import Crypto.Compilers.Z.Named.RewriteAddToAdc.
-Require Import Crypto.Util.PointedProp.
-Require Import Crypto.Util.Decidable.
-
-(** N.B. This procedure only works when there are no nested lets,
- i.e., nothing like [let x := let y := z in w] in the PHOAS syntax
- tree. This is a limitation of [compile]. *)
-
-Local Open Scope bool_scope.
-
-Section language.
- Local Notation PContext var := (PositiveContext _ var _ internal_base_type_dec_bl).
-
- Definition RewriteAdc {t} (e : Expr t)
- : Expr t
- := let is_good e' := match option_map (wf_unit (Context:=PContext _) empty) e' with
- | Some (Some trivial) => true
- | _ => false
- end in
- let interp_to_phoas := InterpToPHOAS (Context:=fun var => PContext var)
- (fun _ t => Op (OpConst 0%Z) TT) in
- let e' := compile (e _) (DefaultNamesFor e) in
- let e' := option_map (rewrite_expr Pos.eqb) e' in
- let good := is_good e' in
- let e' := match e' with
- | Some e'
- => let ls := Named.default_names_for e' in
- match EliminateDeadCode (Context:=PContext _) e' ls with
- | Some e'' => Some e''
- | None => Some e'
- end
- | None => None
- end in
- let good := good && is_good e' in
- if good
- then let e' := option_map interp_to_phoas e' in
- match e' with
- | Some e'
- => match t return Expr (Arrow (domain t) (codomain t)) -> Expr t with
- | Arrow _ _ => fun x => x
- end e'
- | None => e
- end
- else e.
-End language.
diff --git a/src/Compilers/Z/RewriteAddToAdcInterp.v b/src/Compilers/Z/RewriteAddToAdcInterp.v
deleted file mode 100644
index 28953a3fb..000000000
--- a/src/Compilers/Z/RewriteAddToAdcInterp.v
+++ /dev/null
@@ -1,71 +0,0 @@
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.PositiveContext.DefaultsProperties.
-Require Import Crypto.Compilers.Named.ContextDefinitions.
-Require Import Crypto.Compilers.Named.InterpretToPHOASInterp.
-Require Import Crypto.Compilers.Named.CompileWf.
-Require Import Crypto.Compilers.Named.CompileInterp.
-Require Import Crypto.Compilers.Named.WfFromUnit.
-Require Import Crypto.Compilers.Named.DeadCodeEliminationInterp.
-Require Import Crypto.Compilers.Named.WfInterp.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.RewriteAddToAdc.
-Require Import Crypto.Compilers.Z.Named.RewriteAddToAdcInterp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Bool.
-
-Section language.
- Local Notation PContext var := (PositiveContext _ var _ internal_base_type_dec_bl).
-
- Lemma InterpRewriteAdc
- {t} (e : Expr t) (Hwf : Wf e)
- : forall x, Compilers.Syntax.Interp interp_op (RewriteAdc e) x = Compilers.Syntax.Interp interp_op e x.
- Proof.
- intro x; unfold RewriteAdc, option_map; break_innermost_match; try reflexivity;
- match goal with |- ?x = ?y => cut (Some x = Some y); [ congruence | ] end;
- (etransitivity; [ symmetry; eapply @Interp_InterpToPHOAS with (t:=Arrow _ _) | ]);
- repeat
- repeat
- first [ lazymatch goal with
- | [ H : DeadCodeElimination.EliminateDeadCode _ _ = Some ?e |- Syntax.Named.Interp ?e _ = Some _ ]
- => let lhs := match goal with |- ?lhs = _ => lhs end in
- let v := fresh in
- (destruct lhs as [v|] eqn:?);
- [ apply f_equal; eapply @InterpEliminateDeadCode with (Name_beq:=BinPos.Pos.eqb);
- [ .. | eassumption | try eassumption | try eassumption ]; clear H | ]
- | [ |- Syntax.Named.Interp (RewriteAddToAdc.rewrite_expr _ ?e) _ = Some _ ]
- => let lhs := match goal with |- ?lhs = _ => lhs end in
- let H := fresh in
- destruct lhs eqn:H; [ apply (f_equal (@Some _)); eapply @Interp_rewrite_expr in H | ]
- | [ H : Compile.compile (?e _) _ = Some ?e'', H' : Syntax.Named.Interp ?e'' ?x = Some ?v' |- ?v' = Compilers.Syntax.Interp ?interp_op' ?e ?x ]
- => eapply @Interp_compile with (v:=x) (interp_op:=interp_op') in H
- end
- | intros; exact (@PositiveContextOk _ _ base_type_beq internal_base_type_dec_bl internal_base_type_dec_lb)
- | progress split_andb
- | congruence
- | tauto
- | solve [ auto | eapply @BinPos.Pos.eqb_eq; auto ]
- | eapply @Wf_from_unit
- | eapply @dec_rel_of_bool_dec_rel
- | eapply @internal_base_type_dec_lb
- | eapply @internal_base_type_dec_bl
- | eapply @InterpEliminateDeadCode; [ .. | eassumption | eassumption | ]
- | apply name_list_unique_DefaultNamesFor
- | progress intros
- | rewrite !@lookupb_empty
- | eapply @wf_from_unit with (uContext:=PContext _); [ .. | eassumption ]
- | match goal with
- | [ H : Syntax.Named.Interp ?e ?x = Some ?a, H' : Syntax.Named.Interp ?e ?x = Some ?b |- _ ]
- => assert (a = b) by congruence; (subst a || subst b)
- end
- | lazymatch goal with
- | [ |- Some _ = Some _ ] => fail
- | [ |- None = Some _ ] => exfalso; eapply @wf_interp_not_None; [ .. | unfold Syntax.Named.Interp in *; eassumption ]
- | [ |- ?x = Some _ ] => destruct x eqn:?; [ apply f_equal | ]
- end ].
- Qed.
-End language.
-
-Hint Rewrite @InterpRewriteAdc using solve_wf_side_condition : reflective_interp.
diff --git a/src/Compilers/Z/RewriteAddToAdcWf.v b/src/Compilers/Z/RewriteAddToAdcWf.v
deleted file mode 100644
index 984d2a3e1..000000000
--- a/src/Compilers/Z/RewriteAddToAdcWf.v
+++ /dev/null
@@ -1,39 +0,0 @@
-Require Import Crypto.Compilers.Named.PositiveContext.
-Require Import Crypto.Compilers.Named.InterpretToPHOASWf.
-Require Import Crypto.Compilers.Named.CompileWf.
-Require Import Crypto.Compilers.Named.WfFromUnit.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.RewriteAddToAdc.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.Bool.
-
-Section language.
- Local Hint Resolve internal_base_type_dec_lb internal_base_type_dec_lb dec_rel_of_bool_dec_rel : typeclass_instances.
-
- Lemma Wf_RewriteAdc {t} (e : Expr t) (Hwf : Wf e)
- : Wf (RewriteAdc e).
- Proof.
- unfold RewriteAdc, option_map; break_innermost_match;
- [ .. | solve [ intros var1 var2; destruct (Hwf var1 var2); auto with wf ] ];
- repeat first [ eapply @Wf_InterpToPHOAS with (t:=Arrow _ _)
- | progress split_andb
- | congruence
- | intros;
- match goal with
- | [ |- ContextDefinitions.ContextOk _ ]
- => eapply @PositiveContextOk
- end
- | solve [ auto | eapply @BinPos.Pos.eqb_eq ]
- | eapply @Wf_from_unit
- | eapply @dec_rel_of_bool_dec_rel
- | eapply @internal_base_type_dec_lb
- | eapply @internal_base_type_dec_bl
- | intros var1 var2; specialize (Hwf var1 var2); destruct Hwf;
- constructor; assumption ].
- Qed.
-End language.
-
-Hint Resolve Wf_RewriteAdc : wf.
diff --git a/src/Compilers/Z/Syntax.v b/src/Compilers/Z/Syntax.v
deleted file mode 100644
index dbc739dcb..000000000
--- a/src/Compilers/Z/Syntax.v
+++ /dev/null
@@ -1,103 +0,0 @@
-(** * PHOAS Syntax for expression trees on ℤ *)
-Require Import Coq.ZArith.ZArith.
-Require Import bbv.WordScope.
-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.ZUtil.Definitions.
-Require Import Crypto.Util.IdfunWithAlt.
-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)
-| IdWithAlt T1 T2 Tout : op (Tbase T1 * Tbase T2) (Tbase Tout)
-| Zselect T1 T2 T3 Tout : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout)
-| MulSplit (bitwidth : Z) T1 T2 Tout1 Tout2 : op (Tbase T1 * Tbase T2) (Tbase Tout1 * Tbase Tout2)
-| AddWithCarry T1 T2 T3 Tout : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout)
-| AddWithGetCarry (bitwidth : Z) T1 T2 T3 Tout1 Tout2 : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout1 * Tbase Tout2)
-| SubWithBorrow T1 T2 T3 Tout : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout)
-| SubWithGetBorrow (bitwidth : Z) T1 T2 T3 Tout1 Tout2 : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout1 * Tbase Tout2)
-.
-
-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
- | IdWithAlt _ _ _ => fun xy => id_with_alt (fst xy) (snd xy)
- | Zselect _ _ _ _ => fun ctf => let '(c, t, f) := eta3 ctf in Z.zselect c t f
- | MulSplit bitwidth _ _ _ _ => fun xy => Z.mul_split_at_bitwidth bitwidth (fst xy) (snd xy)
- | AddWithCarry _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.add_with_carry c x y
- | AddWithGetCarry bitwidth _ _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.add_with_get_carry bitwidth c x y
- | SubWithBorrow _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.sub_with_borrow c x y
- | SubWithGetBorrow bitwidth _ _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.sub_with_get_borrow bitwidth c x y
- 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).
-
-Notation Expr := (Expr base_type op).
-Notation Interp := (Interp interp_op).
diff --git a/src/Compilers/Z/Syntax/Equality.v b/src/Compilers/Z/Syntax/Equality.v
deleted file mode 100644
index 0ff0fabf0..000000000
--- a/src/Compilers/Z/Syntax/Equality.v
+++ /dev/null
@@ -1,215 +0,0 @@
-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)) := _.
-
-Notation base_type_dec_bl := internal_base_type_dec_bl.
-Notation base_type_dec_lb := internal_base_type_dec_lb.
-Notation flat_type_beq := (@flat_type_beq base_type base_type_beq).
-Notation flat_type_dec_bl := (@flat_type_dec_bl base_type base_type_beq base_type_dec_bl).
-Notation flat_type_dec_lb := (@flat_type_dec_lb base_type base_type_beq base_type_dec_lb).
-
-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'
- | IdWithAlt T1 T2 Tout, IdWithAlt T1' T2' Tout'
- => base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq Tout Tout'
- | Zselect T1 T2 T3 Tout, Zselect T1' T2' T3' Tout'
- => base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout Tout'
- | MulSplit bitwidth T1 T2 Tout1 Tout2, MulSplit bitwidth' T1' T2' Tout1' Tout2'
- => Z.eqb bitwidth bitwidth' && base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq Tout1 Tout1' && base_type_beq Tout2 Tout2'
- | AddWithCarry T1 T2 T3 Tout, AddWithCarry T1' T2' T3' Tout'
- => base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout Tout'
- | AddWithGetCarry bitwidth T1 T2 T3 Tout1 Tout2, AddWithGetCarry bitwidth' T1' T2' T3' Tout1' Tout2'
- => Z.eqb bitwidth bitwidth' && base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout1 Tout1' && base_type_beq Tout2 Tout2'
- | SubWithBorrow T1 T2 T3 Tout, SubWithBorrow T1' T2' T3' Tout'
- => base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout Tout'
- | SubWithGetBorrow bitwidth T1 T2 T3 Tout1 Tout2, SubWithGetBorrow bitwidth' T1' T2' T3' Tout1' Tout2'
- => Z.eqb bitwidth bitwidth' && base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout1 Tout1' && base_type_beq Tout2 Tout2'
- | OpConst _ _, _
- | Add _ _ _, _
- | Sub _ _ _, _
- | Mul _ _ _, _
- | Shl _ _ _, _
- | Shr _ _ _, _
- | Land _ _ _, _
- | Lor _ _ _, _
- | Opp _ _, _
- | IdWithAlt _ _ _, _
- | Zselect _ _ _ _, _
- | MulSplit _ _ _ _ _, _
- | AddWithCarry _ _ _ _, _
- | AddWithGetCarry _ _ _ _ _ _, _
- | SubWithBorrow _ _ _ _, _
- | SubWithGetBorrow _ _ _ _ _ _, _
- => 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.
-Ltac inversion_base_type_constr_step :=
- lazymatch goal with
- | [ H : TWord _ = TWord _ |- _ ]
- => induction_type_in_using H @path_base_type_rect
- | [ H : TWord _ = TZ |- _ ]
- => induction_type_in_using H @path_base_type_rect
- | [ H : TZ = TWord _ |- _ ]
- => induction_type_in_using H @path_base_type_rect
- | [ H : TZ = TZ |- _ ]
- => induction_type_in_using H @path_base_type_rect
- end.
-Ltac inversion_base_type_constr := repeat inversion_base_type_constr_step.
diff --git a/src/Compilers/Z/Syntax/Util.v b/src/Compilers/Z/Syntax/Util.v
deleted file mode 100644
index 110e6b816..000000000
--- a/src/Compilers/Z/Syntax/Util.v
+++ /dev/null
@@ -1,317 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.micromega.Lia.
-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.HProp.
-Require Import Crypto.Util.Tactics.BreakMatch.
-Require Import Crypto.Util.Tactics.DestructHead.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.PointedProp.
-
-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 is_opp s d (v : op s d) : bool
- := match v with Opp _ _ => true | _ => false end.
-Arguments is_opp [s d] v.
-Definition is_const_or_opp s d (v : op s d) : bool
- := (is_const v || is_opp v)%bool.
-Arguments is_const_or_opp [s d] v.
-
-
-Definition interped_op_side_conditions {s d} (opc : op s d)
- : interp_flat_type interp_base_type s -> pointed_Prop
- := match opc in op s d return interp_flat_type _ s -> _ with
- | IdWithAlt TZ TZ TZ
- => fun v1v2 : Z * Z
- => inject (fst v1v2 = snd v1v2)
- | _ => fun _ => trivial
- end.
-
-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 _ _
- | IdWithAlt _ _ _ => fun _ _ => IdWithAlt _ _ _
- | Zselect _ _ _ _ => fun _ _ => Zselect _ _ _ _
- | MulSplit bitwidth _ _ _ _ => fun _ _ => MulSplit bitwidth _ _ _ _
- | AddWithCarry _ _ _ _ => fun _ _ => AddWithCarry _ _ _ _
- | AddWithGetCarry bitwidth _ _ _ _ _ => fun _ _ => AddWithGetCarry bitwidth _ _ _ _ _
- | SubWithBorrow _ _ _ _ => fun _ _ => SubWithBorrow _ _ _ _
- | SubWithGetBorrow bitwidth _ _ _ _ _ => fun _ _ => SubWithGetBorrow bitwidth _ _ _ _ _
- 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_small {a b c} v
- : match b with
- | TZ => True
- | TWord bsz => 0 <= interpToZ (@cast_const a c v) < 2^Z.of_nat (2^bsz)
- end%Z
- -> @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 lia
- | rewrite wordToZ_ZToWord_wordToZ by lia
- | rewrite wordToZ_ZToWord by assumption
- | rewrite ZToWord_wordToZ_ZToWord_small by omega ].
-Qed.
-
-Lemma cast_const_split_mod {a b} v
- : @cast_const a b v = ZToInterp (match a, b with
- | TZ, _ => interpToZ v
- | _, TWord lgsz => (interpToZ v) mod (2^Z.of_nat (2^lgsz))
- | _, TZ => interpToZ v
- end).
-Proof.
- destruct_head base_type; simpl; try reflexivity.
- rewrite <- wordToZ_ZToWord_mod, ZToWord_wordToZ by apply wordToZ_range.
- reflexivity.
-Qed.
-
-Lemma interpToZ_cast_const_mod {a b} v
- : interpToZ (@cast_const a b v)
- = match b with
- | TZ => interpToZ v
- | TWord lgsz => Z.max 0 (interpToZ v) mod (2^Z.of_nat (2^lgsz))
- end%Z.
-Proof.
- repeat first [ progress destruct_head base_type
- | reflexivity
- | rewrite wordToZ_ZToWord_mod_full ].
-Qed.
-
-Lemma cast_const_ZToInterp_mod {a b} v
- : @cast_const a b (ZToInterp v)
- = ZToInterp match a with
- | TZ => v
- | TWord lgsz => Z.max 0 v mod 2^Z.of_nat (2^lgsz)
- end%Z.
-Proof.
- repeat first [ progress destruct_head base_type
- | reflexivity
- | rewrite wordToZ_ZToWord_mod_full ].
-Qed.
-
-Lemma interpToZ_ZToInterp_mod {a} v
- : @interpToZ a (ZToInterp v)
- = match a with
- | TZ => v
- | TWord lgsz => Z.max 0 v mod 2^Z.of_nat (2^lgsz)
- end%Z.
-Proof.
- etransitivity; [ apply (@interpToZ_cast_const_mod TZ) | ].
- 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 lia
- | rewrite wordToZ_ZToWord_wordToZ by lia ].
-Qed.
-
-Lemma ZToInterp_eq_inj {a} x y
- : @ZToInterp a x = @ZToInterp a y
- <-> match a with
- | TZ => x
- | TWord lgsz => Z.max 0 x mod 2^Z.of_nat (2^lgsz)
- end%Z
- = match a with
- | TZ => y
- | TWord lgsz => Z.max 0 y mod 2^Z.of_nat (2^lgsz)
- end%Z.
-Proof.
- rewrite <- !interpToZ_ZToInterp_mod.
- destruct a; try reflexivity; simpl.
- split; intro H; try congruence.
- rewrite <- (ZToWord_wordToZ (FixedWordSizes.ZToWord x)), <- (ZToWord_wordToZ (FixedWordSizes.ZToWord y)).
- congruence.
-Qed.
-
-Lemma interpToZ_range {a} (v : interp_base_type a)
- : match a with
- | TZ => True
- | TWord lgsz => 0 <= interpToZ v < 2^Z.of_nat (2^lgsz)
- end%Z.
-Proof.
- destruct a; trivial; simpl.
- apply wordToZ_range.
-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.
-
-Lemma base_type_leb_total
- : forall x y : base_type, base_type_leb x y = true \/ base_type_leb y x = true.
-Proof.
- induction x, y; simpl; auto.
- rewrite !Nat.leb_le; omega.
-Qed.
-
-Lemma eta_match_base_type_impl P1 P2 PZ T
- : match T as T return P1 T -> P2 with
- | TZ => fun _ => PZ
- | TWord _ => fun _ => PZ
- end = fun _ => PZ.
-Proof. destruct T; reflexivity. Qed.
-Ltac rewrite_eta_match_base_type_impl_step :=
- match goal with
- | [ H : context[match ?T as T' in base_type return (@?P1 T' -> ?P2) with TZ => fun _ => ?PZ | TWord _ => fun _ => ?PZ end] |- _ ]
- => rewrite (@eta_match_base_type_impl P1 P2 PZ T) in H
- | [ |- context[match ?T as T' in base_type return (@?P1 T' -> ?P2) with TZ => fun _ => ?PZ | TWord _ => fun _ => ?PZ end] ]
- => rewrite (@eta_match_base_type_impl P1 P2 PZ T)
- end.
-Ltac rewrite_eta_match_base_type_impl := repeat rewrite_eta_match_base_type_impl_step.
diff --git a/src/Compilers/Z/TypeInversion.v b/src/Compilers/Z/TypeInversion.v
deleted file mode 100644
index 92338dcfb..000000000
--- a/src/Compilers/Z/TypeInversion.v
+++ /dev/null
@@ -1,43 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.TypeInversion.
-Require Import Crypto.Util.FixCoqMistakes.
-
-Section language.
- Section flat.
- Context (P : flat_type base_type -> Type).
-
- Local Ltac t :=
- let H := fresh in
- intro H; intros;
- match goal with
- | [ p : _ |- _ ] => specialize (H _ p)
- end;
- cbv beta iota in *;
- try specialize (H eq_refl); simpl in *;
- try assumption.
-
- Definition preinvert_TbaseTZ (Q : P (Tbase TZ) -> Type)
- : (forall t (p : P t), match t return P t -> Type with Tbase TZ => Q | _ => fun _ => True end p)
- -> forall p, Q p.
- Proof. t. Defined.
-
- Definition preinvert_TbaseTWord (Q : forall t, P (Tbase (TWord t)) -> Type)
- : (forall t (p : P t), match t return P t -> Type with Tbase (TWord _) => Q _ | _ => fun _ => True end p)
- -> forall t p, Q t p.
- Proof. t. Defined.
- End flat.
-End language.
-
-Ltac preinvert_one_type e :=
- lazymatch type of e with
- | ?P (Tbase TZ)
- => revert dependent e;
- refine (preinvert_TbaseTZ P _ _)
- | ?P (Tbase (TWord ?T))
- => is_var T;
- move e at top;
- revert dependent T;
- refine (preinvert_TbaseTWord P _ _)
- | _ => Compilers.TypeInversion.preinvert_one_type e
- end.
diff --git a/src/Compilers/ZExtended/InlineConstAndOp.v b/src/Compilers/ZExtended/InlineConstAndOp.v
deleted file mode 100644
index c1fd77e9d..000000000
--- a/src/Compilers/ZExtended/InlineConstAndOp.v
+++ /dev/null
@@ -1,10 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InlineConstAndOp.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-
-Definition inline_const_and_opf {var} {t} (e : exprf _ _ t) : @exprf base_type op var t
- := @inline_const_and_opf base_type op interp_base_type (@interp_op) var (@Const) t e.
-Definition inline_const_and_op {var} {t} (e : expr _ _ t) : @expr base_type op var t
- := @inline_const_and_op base_type op interp_base_type (@interp_op) var (@Const) t e.
-Definition InlineConstAndOp {t} (e : Expr t) : Expr t
- := @InlineConstAndOp base_type op interp_base_type (@interp_op) (@Const) t e.
diff --git a/src/Compilers/ZExtended/InlineConstAndOpByRewrite.v b/src/Compilers/ZExtended/InlineConstAndOpByRewrite.v
deleted file mode 100644
index 418a0e0f3..000000000
--- a/src/Compilers/ZExtended/InlineConstAndOpByRewrite.v
+++ /dev/null
@@ -1,12 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InlineConstAndOpByRewrite.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-
-Module Export Rewrite.
- Definition inline_const_and_opf {var} {t} (e : exprf _ _ t) : @exprf base_type op var t
- := @inline_const_and_opf base_type op interp_base_type (@interp_op) var (@Const) t e.
- Definition inline_const_and_op {var} {t} (e : expr _ _ t) : @expr base_type op var t
- := @inline_const_and_op base_type op interp_base_type (@interp_op) var (@Const) t e.
- Definition InlineConstAndOp {t} (e : Expr t) : Expr t
- := @InlineConstAndOp base_type op interp_base_type (@interp_op) (@Const) t e.
-End Rewrite.
diff --git a/src/Compilers/ZExtended/InlineConstAndOpByRewriteInterp.v b/src/Compilers/ZExtended/InlineConstAndOpByRewriteInterp.v
deleted file mode 100644
index 64d5263a7..000000000
--- a/src/Compilers/ZExtended/InlineConstAndOpByRewriteInterp.v
+++ /dev/null
@@ -1,15 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.InlineConstAndOpByRewriteInterp.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-Require Import Crypto.Compilers.ZExtended.InlineConstAndOpByRewrite.
-
-Module Export Rewrite.
- Lemma InterpInlineConstAndOp {t} (e : Expr t)
- : forall x, Interp (InlineConstAndOp e) x = Interp e x.
- Proof.
- refine (@InterpInlineConstAndOp _ _ _ _ _ t e _).
- clear; abstract (intros []; intros; reflexivity).
- Defined.
-
- Hint Rewrite @InterpInlineConstAndOp : reflective_interp.
-End Rewrite.
diff --git a/src/Compilers/ZExtended/InlineConstAndOpByRewriteWf.v b/src/Compilers/ZExtended/InlineConstAndOpByRewriteWf.v
deleted file mode 100644
index 281bee6ab..000000000
--- a/src/Compilers/ZExtended/InlineConstAndOpByRewriteWf.v
+++ /dev/null
@@ -1,13 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.InlineConstAndOpByRewriteWf.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-Require Import Crypto.Compilers.ZExtended.InlineConstAndOpByRewrite.
-
-Module Export Rewrite.
- Definition Wf_InlineConstAndOp {t} (e : Expr t) (Hwf : Wf e)
- : Wf (InlineConstAndOp e)
- := @Wf_InlineConstAndOp _ _ _ _ _ t e Hwf.
-
- Hint Resolve Wf_InlineConstAndOp : wf.
-End Rewrite.
diff --git a/src/Compilers/ZExtended/InlineConstAndOpInterp.v b/src/Compilers/ZExtended/InlineConstAndOpInterp.v
deleted file mode 100644
index 3f1a83f3d..000000000
--- a/src/Compilers/ZExtended/InlineConstAndOpInterp.v
+++ /dev/null
@@ -1,14 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.InlineConstAndOpInterp.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-Require Import Crypto.Compilers.ZExtended.InlineConstAndOp.
-
-Definition InterpInlineConstAndOp {t} (e : Expr t) (Hwf : Wf e)
- : forall x, Interp (InlineConstAndOp e) x = Interp e x.
-Proof.
- refine (@InterpInlineConstAndOp _ _ _ _ _ t e Hwf _).
- clear; abstract (intros []; intros; reflexivity).
-Defined.
-
-Hint Rewrite @InterpInlineConstAndOp using solve_wf_side_condition : reflective_interp.
diff --git a/src/Compilers/ZExtended/InlineConstAndOpWf.v b/src/Compilers/ZExtended/InlineConstAndOpWf.v
deleted file mode 100644
index 3ce24e237..000000000
--- a/src/Compilers/ZExtended/InlineConstAndOpWf.v
+++ /dev/null
@@ -1,11 +0,0 @@
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Wf.
-Require Import Crypto.Compilers.InlineConstAndOpWf.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-Require Import Crypto.Compilers.ZExtended.InlineConstAndOp.
-
-Definition Wf_InlineConstAndOp {t} (e : Expr t) (Hwf : Wf e)
- : Wf (InlineConstAndOp e)
- := @Wf_InlineConstAndOp _ _ _ _ _ t e Hwf.
-
-Hint Resolve Wf_InlineConstAndOp : wf.
diff --git a/src/Compilers/ZExtended/MapBaseType.v b/src/Compilers/ZExtended/MapBaseType.v
deleted file mode 100644
index 66b18f6f5..000000000
--- a/src/Compilers/ZExtended/MapBaseType.v
+++ /dev/null
@@ -1,22 +0,0 @@
-Require Import Coq.ZArith.BinIntDef.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.Util.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-Require Import Crypto.Compilers.ZExtended.Syntax.Util.
-Require Import Crypto.Compilers.MapBaseType.
-
-Definition MapBaseType' {t} (e : Expr t)
- : Z.Syntax.Expr (Arrow (lift_flat_type unextend_base_type (domain t))
- (lift_flat_type unextend_base_type (codomain t)))
- := @MapBaseType' _ _ _ _ unextend_base_type (fun _ s d opc _ => unextend_op opc)
- (fun _ t => Op (make_const t (ZToInterp 0)) TT)
- t e.
-
-Definition MapBaseType {t} (e : Expr t)
- : option (Z.Syntax.Expr (Arrow (lift_flat_type unextend_base_type (domain t))
- (lift_flat_type unextend_base_type (codomain t))))
- := @MapBaseType _ _ _ _ unextend_base_type (fun _ s d opc _ => unextend_op opc)
- (fun _ t => Op (make_const t (ZToInterp 0)) TT)
- t e.
diff --git a/src/Compilers/ZExtended/Syntax.v b/src/Compilers/ZExtended/Syntax.v
deleted file mode 100644
index 77e6148e5..000000000
--- a/src/Compilers/ZExtended/Syntax.v
+++ /dev/null
@@ -1,89 +0,0 @@
-(** * PHOAS Syntax for expression trees on ℤ *)
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Util.Curry.
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.TypeUtil.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.ZUtil.Definitions.
-Require Import Crypto.Util.IdfunWithAlt.
-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 | TBool.
-
-Local Notation tZ := (Tbase TZ).
-Local Notation tBool := (Tbase TBool).
-
-Definition interp_base_type (v : base_type) : Set :=
- match v with
- | TZ => Z
- | TBool => bool
- end.
-
-Inductive op : flat_type base_type -> flat_type base_type -> Set :=
-| AddWithGetCarry : op (tuple tZ 4) (tuple tZ 2)
-| SubWithGetBorrow : op (tuple tZ 4) (tuple tZ 2)
-| MulSplitAtBitwidth : op (tuple tZ 3) (tuple tZ 2)
-| AddWithGetCarryZ (bitwidth : Z) : op (tuple tZ 3) (tuple tZ 2)
-| SubWithGetBorrowZ (bitwidth : Z) : op (tuple tZ 3) (tuple tZ 2)
-| MulSplitAtBitwidthZ (bitwidth : Z) : op (tuple tZ 2) (tuple tZ 2)
-| IdWithAlt {T} : op (tuple T 2) (tuple T 1)
-| Zselect : op (tuple tZ 3) (tuple tZ 1)
-| Zmul : op (tuple tZ 2) (tuple tZ 1)
-| Zadd : op (tuple tZ 2) (tuple tZ 1)
-| Zsub : op (tuple tZ 2) (tuple tZ 1)
-| Zopp : op (tuple tZ 1) (tuple tZ 1)
-| Zshiftr : op (tuple tZ 2) (tuple tZ 1)
-| Zshiftl : op (tuple tZ 2) (tuple tZ 1)
-| Zland : op (tuple tZ 2) (tuple tZ 1)
-| Zlor : op (tuple tZ 2) (tuple tZ 1)
-| Zmodulo : op (tuple tZ 2) (tuple tZ 1)
-| Zdiv : op (tuple tZ 2) (tuple tZ 1)
-| Zlog2 : op (tuple tZ 1) (tuple tZ 1)
-| Zpow : op (tuple tZ 2) (tuple tZ 1)
-| Zones : op (tuple tZ 1) (tuple tZ 1)
-| Zeqb : op (tuple tZ 2) (tuple tBool 1)
-| ConstZ (v : BinNums.Z) : op (tuple tZ 0) (tuple tZ 1)
-| ConstBool (v : bool) : op (tuple tZ 0) (tuple tBool 1)
-| BoolCase {T} : op (Prod (Prod tBool T) T) T.
-
-Definition Const {t} : interp_base_type t -> op Unit (Tbase t)
- := match t with
- | TZ => ConstZ
- | Tbool => ConstBool
- end.
-
-Definition interp_op {s d} (opv : op s d) : interp_flat_type interp_base_type s -> interp_flat_type interp_base_type d
- := match opv with
- | AddWithGetCarry => curry4 Z.add_with_get_carry
- | SubWithGetBorrow => curry4 Z.sub_with_get_borrow
- | MulSplitAtBitwidth => curry3 Z.mul_split_at_bitwidth
- | AddWithGetCarryZ bitwidth => curry3 (Z.add_with_get_carry bitwidth)
- | SubWithGetBorrowZ bitwidth => curry3 (Z.sub_with_get_borrow bitwidth)
- | MulSplitAtBitwidthZ bitwidth => curry2 (Z.mul_split_at_bitwidth bitwidth)
- | IdWithAlt T => curry2 id_with_alt
- | Zselect => curry3 Z.zselect
- | Zmul => curry2 Z.mul
- | Zadd => curry2 Z.add
- | Zsub => curry2 Z.sub
- | Zopp => Z.opp
- | Zshiftr => curry2 Z.shiftr
- | Zshiftl => curry2 Z.shiftl
- | Zland => curry2 Z.land
- | Zlor => curry2 Z.lor
- | Zmodulo => curry2 Z.modulo
- | Zdiv => curry2 Z.div
- | Zlog2 => Z.log2
- | Zpow => curry2 Z.pow
- | Zones => Z.ones
- | Zeqb => curry2 Z.eqb
- | ConstZ v => fun _ => v
- | ConstBool v => fun _ => v
- | BoolCase T => fun '(b, t, f) => if b then t else f
- end.
-
-Notation Expr := (Expr base_type op).
-Notation Interp := (Interp (@interp_op)).
diff --git a/src/Compilers/ZExtended/Syntax/Util.v b/src/Compilers/ZExtended/Syntax/Util.v
deleted file mode 100644
index 7c3ed70e6..000000000
--- a/src/Compilers/ZExtended/Syntax/Util.v
+++ /dev/null
@@ -1,81 +0,0 @@
-Require Import Crypto.Compilers.SmartMap.
-Require Import Crypto.Compilers.ExprInversion.
-Require Import Crypto.Compilers.Syntax.
-Require Import Crypto.Compilers.Z.Syntax.
-Require Import Crypto.Compilers.ZExtended.Syntax.
-
-Definition unextend_base_type (t : ZExtended.Syntax.base_type) : Z.Syntax.base_type
- := match t with
- | TZ => Z.Syntax.TZ
- | TBool => Z.Syntax.TZ
- end.
-
-Definition invert_Const s d (opc : op s d) : option (interp_flat_type interp_base_type d)
- := match opc with
- | ConstZ v => Some v
- | ConstBool v => Some v
- | AddWithGetCarry
- | SubWithGetBorrow
- | MulSplitAtBitwidth
- | AddWithGetCarryZ _
- | SubWithGetBorrowZ _
- | MulSplitAtBitwidthZ _
- | IdWithAlt _
- | Zselect
- | Zmul
- | Zadd
- | Zsub
- | Zopp
- | Zshiftr
- | Zshiftl
- | Zland
- | Zlor
- | Zmodulo
- | Zdiv
- | Zlog2
- | Zpow
- | Zones
- | Zeqb
- | BoolCase _
- => None
- end.
-
-Definition unextend_op {s d} (opc : ZExtended.Syntax.op s d)
- : option (Z.Syntax.op (lift_flat_type unextend_base_type s)
- (lift_flat_type unextend_base_type d))
- := match opc in ZExtended.Syntax.op s d
- return option (Z.Syntax.op (lift_flat_type unextend_base_type s)
- (lift_flat_type unextend_base_type d))
- with
- | AddWithGetCarry => None
- | AddWithGetCarryZ bitwidth
- => Some (Z.Syntax.AddWithGetCarry bitwidth _ _ _ _ _)
- | SubWithGetBorrow => None
- | SubWithGetBorrowZ bitwidth
- => Some (Z.Syntax.SubWithGetBorrow bitwidth _ _ _ _ _)
- | MulSplitAtBitwidth => None
- | MulSplitAtBitwidthZ bitwidth
- => Some (Z.Syntax.MulSplit bitwidth _ _ _ _)
- | IdWithAlt (Tbase _)
- => Some (Z.Syntax.IdWithAlt _ _ _)
- | IdWithAlt _
- => None
- | Zselect => Some (Z.Syntax.Zselect _ _ _ _)
- | Zmul => Some (Z.Syntax.Mul _ _ _)
- | Zadd => Some (Z.Syntax.Add _ _ _)
- | Zsub => Some (Z.Syntax.Sub _ _ _)
- | Zopp => Some (Z.Syntax.Opp _ _)
- | Zshiftr => Some (Z.Syntax.Shr _ _ _)
- | Zshiftl => Some (Z.Syntax.Shl _ _ _)
- | Zland => Some (Z.Syntax.Land _ _ _)
- | Zlor => Some (Z.Syntax.Lor _ _ _)
- | Zmodulo => None
- | Zdiv => None
- | Zlog2 => None
- | Zpow => None
- | Zones => None
- | Zeqb => None
- | ConstZ v => Some (OpConst v)
- | ConstBool v => None
- | BoolCase T => None
- end.