aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--Makefile21
-rw-r--r--_CoqProject35
-rw-r--r--src/Reflection/Z/Interpretations128.v13
-rw-r--r--src/Reflection/Z/Interpretations128/Relations.v572
-rw-r--r--src/Reflection/Z/Interpretations128/RelationsCombinations.v253
-rw-r--r--src/Reflection/Z/Interpretations64.v13
-rw-r--r--src/Reflection/Z/Interpretations64/Relations.v572
-rw-r--r--src/Reflection/Z/Interpretations64/RelationsCombinations.v253
-rw-r--r--src/Reflection/Z/InterpretationsGen.v915
-rw-r--r--src/Specific/GF25519BoundedCommon.v823
-rw-r--r--src/Specific/GF25519Reflective/Common.v669
-rw-r--r--src/Specific/GF25519Reflective/Common9_4Op.v102
-rw-r--r--src/Specific/GF25519Reflective/CommonBinOp.v50
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOp.v43
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOpFEToWire.v43
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOpFEToZ.v43
-rw-r--r--src/Specific/GF25519Reflective/CommonUnOpWireToFE.v43
-rw-r--r--src/Specific/GF25519Reflective/Reified.v13
-rw-r--r--src/Specific/GF25519Reflective/Reified/Add.v12
-rw-r--r--src/Specific/GF25519Reflective/Reified/AddCoordinates.v209
-rw-r--r--src/Specific/GF25519Reflective/Reified/AddDisplay.log45
-rw-r--r--src/Specific/GF25519Reflective/Reified/AddDisplay.v4
-rw-r--r--src/Specific/GF25519Reflective/Reified/AddJavaDisplay.log45
-rw-r--r--src/Specific/GF25519Reflective/Reified/AddJavaDisplay.v4
-rw-r--r--src/Specific/GF25519Reflective/Reified/CarryAdd.v17
-rw-r--r--src/Specific/GF25519Reflective/Reified/CarryOpp.v17
-rw-r--r--src/Specific/GF25519Reflective/Reified/CarrySub.v17
-rw-r--r--src/Specific/GF25519Reflective/Reified/GeModulus.v17
-rw-r--r--src/Specific/GF25519Reflective/Reified/LadderStep.v189
-rw-r--r--src/Specific/GF25519Reflective/Reified/LadderStepDisplay.log3459
-rw-r--r--src/Specific/GF25519Reflective/Reified/LadderStepDisplay.v4
-rw-r--r--src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.log3459
-rw-r--r--src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.v4
-rw-r--r--src/Specific/GF25519Reflective/Reified/Mul.v17
-rw-r--r--src/Specific/GF25519Reflective/Reified/MulDisplay.log297
-rw-r--r--src/Specific/GF25519Reflective/Reified/MulDisplay.v4
-rw-r--r--src/Specific/GF25519Reflective/Reified/MulJavaDisplay.log297
-rw-r--r--src/Specific/GF25519Reflective/Reified/MulJavaDisplay.v4
-rw-r--r--src/Specific/GF25519Reflective/Reified/Opp.v12
-rw-r--r--src/Specific/GF25519Reflective/Reified/Pack.v17
-rw-r--r--src/Specific/GF25519Reflective/Reified/PreFreeze.v17
-rw-r--r--src/Specific/GF25519Reflective/Reified/Sub.v12
-rw-r--r--src/Specific/GF25519Reflective/Reified/Unpack.v17
-rwxr-xr-xsrc/Specific/GF25519Reflective/Reified/rebuild-reified.py30
44 files changed, 7 insertions, 12695 deletions
diff --git a/Makefile b/Makefile
index e7cb3db58..8060e4907 100644
--- a/Makefile
+++ b/Makefile
@@ -12,7 +12,6 @@ HIDE := $(if $(VERBOSE),,@)
.PHONY: coq clean update-_CoqProject cleanall install \
install-coqprime clean-coqprime coqprime \
- specific-display display \
specific non-specific
SORT_COQPROJECT = sed 's,[^/]*/,~&,g' | env LC_COLLATE=C sort | sed 's,~,,g'
@@ -48,15 +47,9 @@ COQ_VOFILES := $(filter-out $(UNMADE_VOFILES),$(VOFILES))
LITE_VOFILES := $(filter-out $(HEAVY_VOFILES),$(COQ_VOFILES))
SPECIFIC_VO := $(filter src/Specific/%,$(VOFILES))
NON_SPECIFIC_VO := $(filter-out $(SPECIFIC_VO),$(VO_FILES))
-SPECIFIC_DISPLAY_VO := $(filter src/Specific/%Display.vo,$(VOFILES))
-DISPLAY_VO := $(SPECIFIC_DISPLAY_VO)
-DISPLAY_JAVA_VO := $(filter %JavaDisplay.vo,$(DISPLAY_VO))
-DISPLAY_NON_JAVA_VO := $(filter-out $(DISPLAY_JAVA_VO),$(DISPLAY_VO))
specific: $(SPECIFIC_VO) coqprime
non-specific: $(NON_SPECIFIC_VO) coqprime
-specific-display: $(SPECIFIC_DISPLAY_VO:.vo=.log) coqprime
-display: $(DISPLAY_VO:.vo=.log) coqprime
coq: $(COQ_VOFILES) coqprime
lite: $(LITE_VOFILES) coqprime
@@ -86,13 +79,13 @@ Makefile.coq: Makefile _CoqProject
$(SHOW)'COQ_MAKEFILE -f _CoqProject > $@'
$(HIDE)$(COQBIN)coq_makefile -f _CoqProject | sed s'|^\(-include.*\)$$|ifneq ($$(filter-out $(FAST_TARGETS),$$(MAKECMDGOALS)),)~\1~else~ifeq ($$(MAKECMDGOALS),)~\1~endif~endif|g' | tr '~' '\n' | sed s'/^clean:$$/clean::/g' | sed s'/^Makefile: /Makefile-old: /g' | sed s'/^printenv:$$/printenv::/g' > $@
-$(DISPLAY_NON_JAVA_VO:.vo=.log) : %Display.log : %.vo %Display.v src/Reflection/Z/CNotations.vo
- $(SHOW)"COQC $*Display > $@"
- $(HIDE)$(COQC) $(COQDEBUG) $(COQFLAGS) $*Display.v > $@.tmp && mv -f $@.tmp $@
-
-$(DISPLAY_JAVA_VO:.vo=.log) : %JavaDisplay.log : %.vo %JavaDisplay.v src/Reflection/Z/JavaNotations.vo
- $(SHOW)"COQC $*JavaDisplay > $@"
- $(HIDE)$(COQC) $(COQDEBUG) $(COQFLAGS) $*JavaDisplay.v > $@.tmp && mv -f $@.tmp $@
+#$(DISPLAY_NON_JAVA_VO:.vo=.log) : %Display.log : %.vo %Display.v src/Reflection/Z/CNotations.vo
+# $(SHOW)"COQC $*Display > $@"
+# $(HIDE)$(COQC) $(COQDEBUG) $(COQFLAGS) $*Display.v > $@.tmp && mv -f $@.tmp $@
+#
+#$(DISPLAY_JAVA_VO:.vo=.log) : %JavaDisplay.log : %.vo %JavaDisplay.v src/Reflection/Z/JavaNotations.vo
+# $(SHOW)"COQC $*JavaDisplay > $@"
+# $(HIDE)$(COQC) $(COQDEBUG) $(COQFLAGS) $*JavaDisplay.v > $@.tmp && mv -f $@.tmp $@
clean::
rm -f Makefile.coq
diff --git a/_CoqProject b/_CoqProject
index 83600e507..20a560af2 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -201,9 +201,6 @@ src/Reflection/Z/HexNotationConstants.v
src/Reflection/Z/Inline.v
src/Reflection/Z/InlineInterp.v
src/Reflection/Z/InlineWf.v
-src/Reflection/Z/Interpretations128.v
-src/Reflection/Z/Interpretations64.v
-src/Reflection/Z/InterpretationsGen.v
src/Reflection/Z/JavaNotations.v
src/Reflection/Z/MapCastByDeBruijn.v
src/Reflection/Z/MapCastByDeBruijnInterp.v
@@ -214,10 +211,6 @@ src/Reflection/Z/Syntax.v
src/Reflection/Z/Bounds/Interpretation.v
src/Reflection/Z/Bounds/MapCastByDeBruijn.v
src/Reflection/Z/Bounds/Relax.v
-src/Reflection/Z/Interpretations128/Relations.v
-src/Reflection/Z/Interpretations128/RelationsCombinations.v
-src/Reflection/Z/Interpretations64/Relations.v
-src/Reflection/Z/Interpretations64/RelationsCombinations.v
src/Reflection/Z/Syntax/Equality.v
src/Reflection/Z/Syntax/Util.v
src/Spec/CompleteEdwardsCurve.v
@@ -231,40 +224,12 @@ src/Spec/MxDH.v
src/Spec/WeierstrassCurve.v
src/Specific/GF1305.v
src/Specific/GF25519.v
-src/Specific/GF25519BoundedCommon.v
src/Specific/IntegrationTest.v
src/Specific/NewBaseSystemTest.v
src/Specific/SC25519.v
src/Specific/FancyMachine256/Barrett.v
src/Specific/FancyMachine256/Core.v
src/Specific/FancyMachine256/Montgomery.v
-src/Specific/GF25519Reflective/Common.v
-src/Specific/GF25519Reflective/Common9_4Op.v
-src/Specific/GF25519Reflective/CommonBinOp.v
-src/Specific/GF25519Reflective/CommonUnOp.v
-src/Specific/GF25519Reflective/CommonUnOpFEToWire.v
-src/Specific/GF25519Reflective/CommonUnOpFEToZ.v
-src/Specific/GF25519Reflective/CommonUnOpWireToFE.v
-src/Specific/GF25519Reflective/Reified.v
-src/Specific/GF25519Reflective/Reified/Add.v
-src/Specific/GF25519Reflective/Reified/AddCoordinates.v
-src/Specific/GF25519Reflective/Reified/AddDisplay.v
-src/Specific/GF25519Reflective/Reified/AddJavaDisplay.v
-src/Specific/GF25519Reflective/Reified/CarryAdd.v
-src/Specific/GF25519Reflective/Reified/CarryOpp.v
-src/Specific/GF25519Reflective/Reified/CarrySub.v
-src/Specific/GF25519Reflective/Reified/GeModulus.v
-src/Specific/GF25519Reflective/Reified/LadderStep.v
-src/Specific/GF25519Reflective/Reified/LadderStepDisplay.v
-src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.v
-src/Specific/GF25519Reflective/Reified/Mul.v
-src/Specific/GF25519Reflective/Reified/MulDisplay.v
-src/Specific/GF25519Reflective/Reified/MulJavaDisplay.v
-src/Specific/GF25519Reflective/Reified/Opp.v
-src/Specific/GF25519Reflective/Reified/Pack.v
-src/Specific/GF25519Reflective/Reified/PreFreeze.v
-src/Specific/GF25519Reflective/Reified/Sub.v
-src/Specific/GF25519Reflective/Reified/Unpack.v
src/Tactics/VerdiTactics.v
src/Tactics/Algebra_syntax/Nsatz.v
src/Test/Curve25519SpecTestVectors.v
diff --git a/src/Reflection/Z/Interpretations128.v b/src/Reflection/Z/Interpretations128.v
deleted file mode 100644
index fad3168ef..000000000
--- a/src/Reflection/Z/Interpretations128.v
+++ /dev/null
@@ -1,13 +0,0 @@
-(** * Interpretation of PHOAS syntax for expression trees on ℤ *)
-Require Import Coq.ZArith.ZArith.
-Require Export Crypto.Reflection.Z.InterpretationsGen.
-Export Reflection.Syntax.Notations.
-
-Module BitSize128 <: BitSize.
- Definition bit_width : nat := 128%nat.
- Lemma bit_width_pos : (0 < Z.of_nat bit_width)%Z.
- Proof. simpl; omega. Qed.
-End BitSize128.
-
-Module Interpretations128 := InterpretationsGen BitSize128.
-Include Interpretations128.
diff --git a/src/Reflection/Z/Interpretations128/Relations.v b/src/Reflection/Z/Interpretations128/Relations.v
deleted file mode 100644
index b7d86bfb9..000000000
--- a/src/Reflection/Z/Interpretations128/Relations.v
+++ /dev/null
@@ -1,572 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.micromega.Psatz.
-Require Import Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Tuple.
-Require Import Crypto.Reflection.Z.InterpretationsGen.
-Require Import Crypto.Reflection.Z.Interpretations128.
-Require Import Crypto.ModularArithmetic.ModularBaseSystemListZOperationsProofs.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.
-
-Definition proj_eq_rel {A B} (proj : A -> B) (x : A) (y : B) : Prop
- := proj x = y.
-Definition related'_Z {t} (x : BoundedWordW.BoundedWord) (y : Z.interp_base_type t) : Prop
- := proj_eq_rel (BoundedWordW.to_Z' _) x y.
-Definition related_Z t : BoundedWordW.interp_base_type t -> Z.interp_base_type t -> Prop
- := LiftOption.lift_relation (@related'_Z) t.
-Definition related'_wordW {t} (x : BoundedWordW.BoundedWord) (y : WordW.interp_base_type t) : Prop
- := proj_eq_rel (BoundedWordW.to_wordW' _) x y.
-Definition related_wordW t : BoundedWordW.interp_base_type t -> WordW.interp_base_type t -> Prop
- := LiftOption.lift_relation (@related'_wordW) t.
-Definition related_bounds t : BoundedWordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
- := LiftOption.lift_relation2 (proj_eq_rel BoundedWordW.BoundedWordToBounds) t.
-
-Definition related_wordW_Z t : WordW.interp_base_type t -> Z.interp_base_type t -> Prop
- := proj_eq_rel (WordW.to_Z _).
-
-Definition related'_wordW_bounds : WordW.wordW -> ZBounds.bounds -> Prop
- := fun value b => (0 <= Bounds.lower b /\ Bounds.lower b <= WordW.wordWToZ value <= Bounds.upper b /\ Z.log2 (Bounds.upper b) < Z.of_nat WordW.bit_width)%Z.
-Definition related_wordW_bounds : WordW.wordW -> ZBounds.t -> Prop
- := fun value b => match b with
- | Some b => related'_wordW_bounds value b
- | None => True
- end.
-Definition related_wordW_boundsi (t : base_type) : WordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
- := match t with
- | _ => related_wordW_bounds
- end.
-Definition related_wordW_boundsi' (t : base_type) : ZBounds.bounds -> WordW.interp_base_type t -> Prop
- := match t return ZBounds.bounds -> WordW.interp_base_type t -> Prop with
- | _ => fun x y => related'_wordW_bounds y x
- end.
-
-Local Notation related_op R interp_op1 interp_op2
- := (forall (src dst : flat_type base_type) (op : op src dst)
- (sv1 : interp_flat_type _ src) (sv2 : interp_flat_type _ src),
- interp_flat_type_rel_pointwise R sv1 sv2 ->
- interp_flat_type_rel_pointwise R (interp_op1 _ _ op sv1) (interp_op2 _ _ op sv2))
- (only parsing).
-Local Notation related_const R interp f g
- := (forall (t : base_type) (v : interp t), R t (f t v) (g t v))
- (only parsing).
-
-Local Ltac related_const_t :=
- let v := fresh in
- let t := fresh in
- intros t v; destruct t; intros; simpl in *; hnf; simpl;
- cbv [BoundedWordW.wordWToBoundedWord related'_Z LiftOption.of' related_Z related_wordW related'_wordW proj_eq_rel smart_interp_flat_map BoundedWordW.to_Z' WordW.to_Z BoundedWordW.to_wordW' BoundedWordW.of_wordW BoundedWordW.boundedWordToWordW] in *;
- break_innermost_match; simpl;
- first [ tauto
- | Z.ltb_to_lt;
- pose proof (WordW.wordWToZ_log_bound v);
- try omega ].
-
-Lemma related_Z_const : related_const related_Z WordW.interp_base_type BoundedWordW.of_wordW WordW.to_Z.
-Proof. related_const_t. Qed.
-Lemma related_bounds_const : related_const related_bounds WordW.interp_base_type BoundedWordW.of_wordW ZBounds.of_wordW.
-Proof. related_const_t. Qed.
-Lemma related_wordW_const : related_const related_wordW WordW.interp_base_type BoundedWordW.of_wordW (fun _ x => x).
-Proof. related_const_t. Qed.
-
-Local Ltac related_wordW_op_t_step :=
- first [ exact I
- | reflexivity
- | progress intros
- | progress inversion_option
- | progress ZBounds.inversion_bounds
- | progress subst
- | progress destruct_head' False
- | progress destruct_head' prod
- | progress destruct_head' and
- | progress destruct_head' option
- | progress destruct_head' BoundedWordW.BoundedWord
- | progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds'] in *
- | progress simpl @fst in *
- | progress simpl @snd in *
- | progress simpl @BoundedWord.upper in *
- | progress simpl @BoundedWord.lower in *
- | progress break_match
- | progress break_match_hyps
- | congruence
- | match goal with
- | [ H : ?op _ = Some _ |- _ ]
- => let H' := fresh in
- rename H into H';
- first [ pose proof (@BoundedWordW.t_map1_correct _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map2_correct _ _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map4_correct _ _ _ _ _ _ H') as H; clear H' ];
- simpl in H
- | [ H : ?op _ = None |- _ ]
- => let H' := fresh in
- rename H into H';
- first [ pose proof (@BoundedWordW.t_map1_correct_None _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map2_correct_None _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map4_correct_None _ _ _ _ _ H') as H; clear H' ];
- simpl in H
- end
- | progress cbv [related'_wordW proj_eq_rel BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value] in *
- | match goal with
- | [ H : ?op None _ = Some _ |- _ ] => progress simpl in H
- | [ H : ?op _ None = Some _ |- _ ] => progress simpl in H
- | [ H : ?op (Some _) (Some _) = Some _ |- _ ] => progress simpl in H
- | [ H : ?op (Some _) (Some _) = None |- _ ] => progress simpl in H
- end ].
-Local Ltac related_wordW_op_t := repeat related_wordW_op_t_step.
-
-Lemma related_wordW_t_map1 T opW opB pf
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Tbase T) related_wordW sv1 sv2
- -> @related_wordW T (BoundedWordW.t_map1 opW opB pf sv1) (opW sv2).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_wordW_t_map2 T opW opB pf
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Tbase T) (Tbase T)) related_wordW sv1 sv2
- -> @related_wordW T (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opW (fst sv2) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_wordW_t_map4 T opW opB pf
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Prod (Prod (Tbase T) (Tbase T)) (Tbase T)) (Tbase T)) related_wordW sv1 sv2
- -> @related_wordW T (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
- (opW (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_tuples_None_left
- T' n T interp_base_type'
- (R : forall t, LiftOption.interp_base_type' T t -> interp_base_type' t -> Prop)
- (RNone : forall v, R T' None v)
- (v : interp_flat_type interp_base_type' (tuple (Tbase T') (S n)))
- : interp_flat_type_rel_pointwise
- R
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option None))
- v.
-Proof.
- induction n; simpl; intuition.
-Qed.
-
-Lemma related_tuples_Some_left
- T' n T interp_base_type'
- (R : forall t, T -> interp_base_type' t -> Prop)
- u
- (v : interp_flat_type interp_base_type' (tuple (Tbase T') n))
- : interp_flat_type_rel_pointwise
- R
- (flat_interp_untuple (T:=Tbase T') u)
- v
- <-> interp_flat_type_rel_pointwise
- (LiftOption.lift_relation R)
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option (Some u)))
- v.
-Proof.
- destruct n as [|n]; [ reflexivity | ].
- induction n; [ reflexivity | ].
- simpl in *; rewrite <- IHn; clear IHn.
- reflexivity.
-Qed.
-
-Lemma related_tuples_Some_left_ext
- {T' : base_type} {n T interp_base_type'}
- {R : forall t, T -> interp_base_type' t -> Prop}
- {u v u'}
- (H : Tuple.lift_option (flat_interp_tuple (T:=Tbase T') (n:=n) u) = Some u')
- : interp_flat_type_rel_pointwise
- R
- (flat_interp_untuple (T:=Tbase T') u') v
- <-> interp_flat_type_rel_pointwise
- (LiftOption.lift_relation R)
- u v.
-Proof.
- destruct n as [|n]; [ reflexivity | ].
- induction n.
- { simpl in *; subst; reflexivity. }
- { destruct_head_hnf' prod.
- simpl in H; break_match_hyps; inversion_option; inversion_prod; subst.
- simpl; rewrite <- IHn by eassumption; clear IHn.
- reflexivity. }
-Qed.
-
-Lemma related_tuples_proj_eq_rel_untuple
- {T' : base_type} {n T interp_base_type'}
- {proj : forall t, T -> interp_base_type' t}
- {u : Tuple.tuple _ n} {v : Tuple.tuple _ n}
- : interp_flat_type_rel_pointwise
- (fun t => proj_eq_rel (proj t))
- (flat_interp_untuple (T:=Tbase T') u)
- (flat_interp_untuple (T:=Tbase T') v)
- <-> (Tuple.map (proj _) u = v).
-Proof.
- destruct n as [|n]; [ destruct_head_hnf' unit; simpl; tauto | ].
- induction n; [ reflexivity | ].
- destruct_head_hnf' prod.
- simpl @Tuple.tuple.
- rewrite !Tuple.map_S, path_prod_uncurried_iff, <- prod_iff_and; unfold fst, snd.
- rewrite <- IHn.
- reflexivity.
-Qed.
-
-Lemma related_tuples_proj_eq_rel_tuple
- {T' : base_type} {n T interp_base_type'}
- {proj : forall t, T -> interp_base_type' t}
- {u v}
- : interp_flat_type_rel_pointwise
- (fun t => proj_eq_rel (proj t))
- u v
- <-> (Tuple.map (proj _) (flat_interp_tuple (n:=n) (T:=Tbase T') u)
- = flat_interp_tuple (T:=Tbase T') v).
-Proof.
- rewrite <- related_tuples_proj_eq_rel_untuple, !flat_interp_untuple_tuple; reflexivity.
-Qed.
-
-Local Arguments LiftOption.lift_relation2 _ _ _ _ !_ !_ / .
-Lemma related_tuples_lift_relation2_untuple
- (T' : base_type) n T U
- (R : T -> U -> Prop)
- (t : option (Tuple.tuple T (S n)))
- (u : option (Tuple.tuple U (S n)))
- : interp_flat_type_rel_pointwise
- (LiftOption.lift_relation2 R)
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option t))
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option u))
- <-> LiftOption.lift_relation2
- (interp_flat_type_rel_pointwise (fun _ => R))
- T'
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => T) (T:=Tbase T')) t)
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => U) (T:=Tbase T')) u).
-Proof.
- induction n.
- { destruct_head' option; reflexivity. }
- { specialize (IHn (option_map (@fst _ _) t) (option_map (@fst _ _) u)).
- destruct_head' option;
- destruct_head_hnf' prod;
- simpl @option_map in *;
- simpl @LiftOption.lift_relation2 in *;
- try (rewrite <- IHn; reflexivity);
- try (simpl @interp_flat_type_rel_pointwise; tauto). }
-Qed.
-
-Lemma related_tuples_lift_relation2_untuple_ext
- (T' : base_type) {n T U}
- {R : T -> U -> Prop}
- {t u}
- (H : (exists v, Tuple.lift_option (n:=S n) (flat_interp_tuple (T:=Tbase T') t) = Some v)
- \/ (exists v, Tuple.lift_option (n:=S n) (flat_interp_tuple (T:=Tbase T') u) = Some v))
- : interp_flat_type_rel_pointwise
- (LiftOption.lift_relation2 R)
- t u
- <-> LiftOption.lift_relation2
- (interp_flat_type_rel_pointwise (fun _ => R))
- T'
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => T) (T:=Tbase T')) (Tuple.lift_option (flat_interp_tuple (T:=Tbase T') t)))
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => U) (T:=Tbase T')) (Tuple.lift_option (flat_interp_tuple (T:=Tbase T') u))).
-Proof.
- induction n.
- { destruct_head_hnf' option; reflexivity. }
- { specialize (IHn (fst t) (fst u)).
- lazymatch type of IHn with
- | ?T -> _ => let H := fresh in assert (H : T); [ | specialize (IHn H); clear H ]
- end.
- { destruct_head' or; [ left | right ]; destruct_head' ex; destruct_head_hnf' prod; eexists;
- (etransitivity;
- [ | first [ refine (f_equal (option_map (@fst _ _)) (_ : _ = Some (_, _))); eassumption
- | refine (f_equal (option_map (@snd _ _)) (_ : _ = Some (_, _))); eassumption ] ]);
- instantiate; simpl in *; break_match; simpl in *; congruence. }
- destruct_head_hnf' prod;
- destruct_head_hnf' option;
- simpl @fst in *; simpl @snd in *;
- (etransitivity; [ simpl @interp_flat_type_rel_pointwise | reflexivity ]);
- try solve [ repeat first [ progress simpl in *
- | tauto
- | congruence
- | progress destruct_head ex
- | progress destruct_head or
- | progress break_match ] ]. }
-Qed.
-
-Lemma lift_option_flat_interp_tuple
- (T' : base_type) {n T x y}
- : (Tuple.lift_option (n:=S n) (A:=T) (flat_interp_tuple (interp_base_type:=LiftOption.interp_base_type' _) (T:=Tbase T') x) = Some y)
- <-> (x = flat_interp_untuple (T:=Tbase T') (n:=S n) (Tuple.push_option (n:=S n) (Some y))).
-Proof.
- rewrite Tuple.push_lift_option; generalize (Tuple.push_option (Some y)); intro.
- split; intro; subst;
- rewrite ?flat_interp_tuple_untuple, ?flat_interp_untuple_tuple;
- reflexivity.
-Qed.
-
-Lemma lift_option_None_interp_flat_type_rel_pointwise_1
- (T' : base_type) T U n R x y
- (H : interp_flat_type_rel_pointwise (LiftOption.lift_relation2 R) x y)
- (HNone : Tuple.lift_option (A:=T) (n:=S n) (flat_interp_tuple (T:=Tbase T') (n:=S n) x) = None)
- : Tuple.lift_option (A:=U) (n:=S n) (flat_interp_tuple (T:=Tbase T') (n:=S n) y) = None.
-Proof.
- induction n; [ | specialize (IHn (fst x) (fst y) (proj1 H)) ];
- repeat first [ progress destruct_head_hnf' False
- | reflexivity
- | progress inversion_option
- | progress simpl in *
- | progress subst
- | progress specialize_by congruence
- | progress destruct_head_hnf' prod
- | progress destruct_head_hnf' and
- | progress destruct_head_hnf' option
- | progress break_match
- | progress break_match_hyps ].
-Qed.
-
-Local Arguments LiftOption.lift_relation _ _ _ _ !_ _ / .
-Local Arguments LiftOption.of' _ _ !_ / .
-Local Arguments BoundedWordW.BoundedWordToBounds !_ / .
-
-Local Ltac unfold_related_const :=
- intros; hnf; simpl;
- cbv [related_wordW LiftOption.lift_relation LiftOption.of' BoundedWordW.wordWToBoundedWord BoundedWordW.SmartBuildBoundedWord BoundedWordW.of_Z BoundedWordW.of_wordW BoundedWordW.wordWToBoundedWord BoundedWordW.to_Z' BoundedWordW.to_wordW'];
- simpl.
-
-Lemma related_wordW_op : related_op related_wordW (@BoundedWordW.interp_op) (@WordW.interp_op).
-Proof.
- (let op := fresh in intros ?? op; destruct op; simpl);
- try first [ apply related_wordW_t_map1
- | apply related_wordW_t_map2
- | apply related_wordW_t_map4
- | unfold_related_const; break_innermost_match; reflexivity
- | exact (fun _ _ x => x) ].
-Qed.
-
-Lemma related_bounds_t_map1 T opW opB pf
- (HN : opB None = None)
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Tbase T) related_bounds sv1 sv2
- -> @related_bounds T (BoundedWordW.t_map1 opW opB pf sv1) (opB sv2).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_bounds_t_map2 T opW opB pf
- (HN0 : forall v, opB None v = None)
- (HN1 : forall v, opB v None = None)
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Tbase T) (Tbase T)) related_bounds sv1 sv2
- -> @related_bounds T (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opB (fst sv2) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_bounds_t_map4 T opW opB pf
- (HN0 : forall x y z, opB None x y z = None)
- (HN1 : forall x y z, opB x None y z = None)
- (HN2 : forall x y z, opB x y None z = None)
- (HN3 : forall x y z, opB x y z None = None)
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Prod (Prod (Tbase T) (Tbase T)) (Tbase T)) (Tbase T)) related_bounds sv1 sv2
- -> @related_bounds T (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
- (opB (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- destruct_head prod.
- intros; destruct_head' prod.
- progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds' proj_eq_rel] in *.
- destruct_head' option; destruct_head_hnf' and; destruct_head_hnf' False; subst;
- try solve [ simpl; rewrite ?HN0, ?HN1, ?HN2, ?HN3; tauto ];
- [].
- related_wordW_op_t.
-Qed.
-
-Local Arguments Tuple.lift_option : simpl never.
-Local Arguments Tuple.push_option : simpl never.
-Local Arguments Tuple.map : simpl never.
-Local Arguments Tuple.map2 : simpl never.
-
-Local Arguments ZBounds.SmartBuildBounds _ _ / .
-
-Local Ltac related_const_op_t :=
- unfold_related_const; break_innermost_match; try reflexivity; hnf; simpl;
- repeat match goal with
- | [ H : andb _ _ = true |- _ ] => apply andb_prop in H
- | _ => progress destruct_head' and
- | _ => progress Z.ltb_to_lt
- | _ => rewrite WordW.wordWToZ_ZToWordW by (simpl @Z.of_nat; omega)
- | [ H : _ |- _ ] => rewrite WordW.wordWToZ_ZToWordW in H by (simpl @Z.of_nat; omega)
- | [ H : (Z.log2 ?x < ?y)%Z |- _ ]
- => unique assert (x < 2^y)%Z by (apply WordW.log2_lt_pow2_alt_proj_r2l; omega)
- | _ => reflexivity
- | _ => omega
- end.
-
-Lemma related_bounds_op : related_op related_bounds (@BoundedWordW.interp_op) (@ZBounds.interp_op).
-Proof.
- let op := fresh in intros ?? op; destruct op; simpl.
- { related_const_op_t. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map1; intros; destruct_head' option; unfold ZBounds.neg; break_match; reflexivity. }
- { apply related_bounds_t_map4; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map4; intros; destruct_head' option; reflexivity. }
- { exact (fun _ _ x => x). }
-Qed.
-
-(* we [match] and [eexact] rather than [eassumption] so that we can backtrack across hypothesis choice, so that we're hypothesis-order-independent *)
-Local Ltac WordW.Rewrites.wordW_util_arith ::=
- solve [ autorewrite with Zshift_to_pow; omega
- | autorewrite with Zshift_to_pow; nia
- | autorewrite with Zshift_to_pow; auto with zarith
- | match goal with
- | [ H : _ |- _ ]
- => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono | eexact H ];
- autorewrite with Zshift_to_pow; auto using Z.mul_le_mono_nonneg with zarith;
- solve [ omega
- | nia
- | etransitivity; [ eapply Z.div_le_mono | eapply Z.div_le_compat_l ];
- auto with zarith ]
- end
- | apply Z.land_nonneg; WordW.Rewrites.wordW_util_arith
- | match goal with
- | [ H : _ |- _ ]
- => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono | eexact H ];
- instantiate; apply Z.min_case_strong; intros;
- first [ etransitivity; [ apply Z.land_upper_bound_l | ]; omega
- | etransitivity; [ apply Z.land_upper_bound_r | ]; omega ]
- end
- | rewrite Z.log2_lor by omega;
- apply Z.max_case_strong; intro;
- match goal with
- | [ H : _ |- _ ]
- => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono; eexact H | assumption ]
- end
- | eapply Z.le_lt_trans; [ eapply Z.log2_le_mono, neg_upperbound | ];
- WordW.Rewrites.wordW_util_arith
- | (progress unfold ModularBaseSystemListZOperations.cmovne, ModularBaseSystemListZOperations.cmovl, ModularBaseSystemListZOperations.neg); break_match;
- WordW.Rewrites.wordW_util_arith ].
-Local Ltac related_Z_op_t_step :=
- first [ progress related_wordW_op_t_step
- | progress cbv [related'_Z proj_eq_rel BoundedWordW.to_Z' BoundedWordW.to_wordW' WordW.to_Z BoundedWordW.boundedWordToWordW BoundedWord.value] in *
- | autorewrite with push_wordWToZ ].
-Local Ltac related_Z_op_t := repeat related_Z_op_t_step.
-
-Local Notation is_bounded_by value lower upper
- := ((0 <= lower /\ lower <= WordW.wordWToZ value <= upper /\ Z.log2 upper < Z.of_nat WordW.bit_width)%Z)
- (only parsing).
-Local Notation is_in_bounds value bounds
- := (is_bounded_by value (Bounds.lower bounds) (Bounds.upper bounds))
- (only parsing).
-
-Lemma related_Z_t_map1 T opZ opW opB pf
- (H : forall x bxs brs,
- Some brs = opB (Some bxs)
- -> is_in_bounds x bxs
- -> is_in_bounds (opW x) brs
- -> WordW.wordWToZ (opW x) = (opZ (WordW.wordWToZ x)))
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Tbase T) related_Z sv1 sv2
- -> @related_Z T (BoundedWordW.t_map1 opW opB pf sv1) (opZ sv2).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_Z_op_t.
- eapply H; eauto.
-Qed.
-
-Lemma related_Z_t_map2 T opZ opW opB pf
- (H : forall x y bxs bys brs,
- Some brs = opB (Some bxs) (Some bys)
- -> is_in_bounds x bxs
- -> is_in_bounds y bys
- -> is_in_bounds (opW x y) brs
- -> WordW.wordWToZ (opW x y) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y)))
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Tbase T) (Tbase T)) related_Z sv1 sv2
- -> @related_Z T (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opZ (fst sv2) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_Z_op_t.
- eapply H; eauto.
-Qed.
-
-Lemma related_Z_t_map4 T opZ opW opB pf
- (H : forall x y z w bxs bys bzs bws brs,
- Some brs = opB (Some bxs) (Some bys) (Some bzs) (Some bws)
- -> is_in_bounds x bxs
- -> is_in_bounds y bys
- -> is_in_bounds z bzs
- -> is_in_bounds w bws
- -> is_in_bounds (opW x y z w) brs
- -> WordW.wordWToZ (opW x y z w) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y) (WordW.wordWToZ z) (WordW.wordWToZ w)))
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=(Tbase T * Tbase T * Tbase T * Tbase T)%ctype) related_Z sv1 sv2
- -> @related_Z T (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
- (opZ (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_Z_op_t.
- eapply H; eauto.
-Qed.
-
-Local Arguments related_Z _ !_ _ / .
-
-Local Arguments related'_Z _ _ _ / .
-
-Local Ltac related_Z_op_fin_t_step :=
- first [ progress subst
- | progress inversion_option
- | progress ZBounds.inversion_bounds
- | progress destruct_head' Bounds.bounds
- | progress destruct_head' ZBounds.bounds
- | progress destruct_head' and
- | progress simpl in * |-
- | progress break_match_hyps
- | congruence
- | progress inversion_option
- | intro
- | progress autorewrite with push_wordWToZ
- | match goal with
- | [ H : andb _ _ = true |- _ ] => rewrite Bool.andb_true_iff in H
- | [ H : context[Tuple.lift_option (Tuple.push_option _)] |- _ ]
- => rewrite Tuple.lift_push_option in H
- end
- | progress Z.ltb_to_lt ].
-Local Ltac related_Z_op_fin_t := repeat related_Z_op_fin_t_step.
-
-Local Opaque WordW.bit_width.
-
-Local Arguments ZBounds.neg _ !_ / .
-
-Lemma related_Z_op : related_op related_Z (@BoundedWordW.interp_op) (@Z.interp_op).
-Proof.
- let op := fresh in intros ?? op; destruct op; simpl.
- { related_const_op_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map1; related_Z_op_fin_t. }
- { apply related_Z_t_map4; related_Z_op_fin_t. }
- { apply related_Z_t_map4; related_Z_op_fin_t. }
- { exact (fun _ _ x => x). }
-Qed.
-
-Create HintDb interp_related discriminated.
-Hint Resolve related_Z_op related_bounds_op related_wordW_op related_Z_const related_bounds_const related_wordW_const : interp_related.
diff --git a/src/Reflection/Z/Interpretations128/RelationsCombinations.v b/src/Reflection/Z/Interpretations128/RelationsCombinations.v
deleted file mode 100644
index b2f7ce4a5..000000000
--- a/src/Reflection/Z/Interpretations128/RelationsCombinations.v
+++ /dev/null
@@ -1,253 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.Classes.RelationClasses.
-Require Import Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Z.InterpretationsGen.
-Require Import Crypto.Reflection.Z.Interpretations128.
-Require Import Crypto.Reflection.Z.Interpretations128.Relations.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.
-
-Module Relations.
- Section proj.
- Context {interp_base_type1 interp_base_type2}
- (proj : forall t : base_type, interp_base_type1 t -> interp_base_type2 t).
-
- Let R {t : flat_type base_type} f g :=
- SmartVarfMap (t:=t) proj f = g.
-
- Definition interp_type_rel_pointwise_uncurried_proj
- {t : type base_type}
- : interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> Prop
- := fun f g
- => forall x : interp_flat_type interp_base_type1 (domain t),
- let y := SmartVarfMap proj x in
- let fx := f x in
- let gy := g y in
- @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj
- {t : type base_type}
- {f}
- {g}
- : (forall x y, @R (domain t) x y -> @R (codomain t) (f x) (g y))
- -> interp_type_rel_pointwise_uncurried_proj (t:=t) f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj.
- destruct t as [A B]; simpl in *.
- subst R; simpl in *.
- eauto.
- Qed.
- End proj.
-
- Section proj_option.
- Context {interp_base_type1 : Type} {interp_base_type2 : base_type -> Type}
- (proj_option : forall t, interp_base_type1 -> interp_base_type2 t).
-
- Let R {t : flat_type base_type} f g :=
- let f' := LiftOption.of' (t:=t) f in
- match f' with
- | Some f' => SmartVarfMap proj_option f' = g
- | None => True
- end.
-
- Definition interp_type_rel_pointwise_uncurried_proj_option
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type interp_base_type2 t -> Prop
- := fun f g
- => forall x : interp_flat_type (fun _ => interp_base_type1) (domain t),
- let y := SmartVarfMap proj_option x in
- let fx := f (LiftOption.to' (Some x)) in
- let gy := g y in
- @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj_option
- {t : type base_type}
- {f}
- {g}
- : (forall x y, @R (domain t) x y -> @R (codomain t) (f x) (g y))
- -> interp_type_rel_pointwise_uncurried_proj_option (t:=t) f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj_option.
- destruct t as [A B]; simpl in *.
- subst R; simpl in *.
- intros H x; apply H; simpl.
- rewrite LiftOption.of'_to'; reflexivity.
- Qed.
- End proj_option.
-
- Section proj_option2.
- Context {interp_base_type1 : Type} {interp_base_type2 : Type}
- (proj : interp_base_type1 -> interp_base_type2).
-
- Let R {t : flat_type base_type} f g :=
- let f' := LiftOption.of' (t:=t) f in
- let g' := LiftOption.of' (t:=t) g in
- match f', g' with
- | Some f', Some g' => SmartVarfMap (fun _ => proj) f' = g'
- | None, None => True
- | Some _, _ => False
- | None, _ => False
- end.
-
- Definition interp_type_rel_pointwise_uncurried_proj_option2
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type (LiftOption.interp_base_type' interp_base_type2) t -> Prop
- := fun f g
- => forall x : interp_flat_type (fun _ => interp_base_type1) (domain t),
- let y := SmartVarfMap (fun _ => proj) x in
- let fx := f (LiftOption.to' (Some x)) in
- let gy := g (LiftOption.to' (Some y)) in
- @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj_option2
- {t : type base_type}
- {f}
- {g}
- : (forall x y, @R (domain t) x y -> @R (codomain t) (f x) (g y))
- -> interp_type_rel_pointwise_uncurried_proj_option2 (t:=t) f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj_option2.
- destruct t as [A B]; simpl in *.
- subst R; simpl in *.
- intros H x; apply H; simpl.
- rewrite !LiftOption.of'_to'; reflexivity.
- Qed.
- End proj_option2.
-
- Local Ltac t proj012 :=
- repeat match goal with
- | _ => progress cbv beta in *
- | _ => progress break_match; try tauto; []
- | _ => progress subst
- | [ H : _ |- _ ]
- => first [ rewrite !LiftOption.lift_relation_flat_type_pointwise in H
- by (eassumption || rewrite LiftOption.of'_to'; reflexivity)
- | rewrite !LiftOption.lift_relation2_flat_type_pointwise in H
- by (eassumption || rewrite LiftOption.of'_to'; reflexivity)
- | rewrite !@lift_interp_flat_type_rel_pointwise_f_eq_id2 in H
- | rewrite !@lift_interp_flat_type_rel_pointwise_f_eq in H ]
- | _ => progress unfold proj_eq_rel in *
- | _ => progress specialize_by reflexivity
- | _ => rewrite SmartVarfMap_compose
- | _ => setoid_rewrite proj012
- | [ |- context G[fun x y => ?f x y] ]
- => let G' := context G[f] in change G'
- | [ |- context G[fun (_ : ?T) x => ?f x] ]
- => let G' := context G[fun _ : T => f] in change G'
- | [ H : context G[fun (_ : ?T) x => ?f x] |- _ ]
- => let G' := context G[fun _ : T => f] in change G' in H
- | _ => rewrite_hyp <- !*; []
- | _ => reflexivity
- | _ => rewrite interp_flat_type_rel_pointwise_SmartVarfMap
- end.
-
- Section proj_from_option2.
- Context {interp_base_type0 : Type} {interp_base_type1 interp_base_type2 : base_type -> Type}
- (proj01 : forall t, interp_base_type0 -> interp_base_type1 t)
- (proj02 : forall t, interp_base_type0 -> interp_base_type2 t)
- (proj : forall t, interp_base_type1 t -> interp_base_type2 t).
-
- Let R {t : flat_type base_type} f g :=
- proj_eq_rel (SmartVarfMap proj (t:=t)) f g.
-
- Definition interp_type_rel_pointwise_uncurried_proj_from_option2
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type0) t -> interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> Prop
- := fun f0 f g
- => forall x : interp_flat_type (fun _ => interp_base_type0) (domain t),
- let x' := SmartVarfMap proj01 x in
- let y' := SmartVarfMap proj x' in
- let fx := f x' in
- let gy := g y' in
- let f0x := LiftOption.of' (f0 (LiftOption.to' (Some x))) in
- match f0x with
- | Some _ => True
- | None => False
- end
- -> @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj_from_option2
- {t : type base_type}
- {f0}
- {f : interp_type interp_base_type1 t}
- {g : interp_type interp_base_type2 t}
- (proj012 : forall t x, proj t (proj01 t x) = proj02 t x)
- : interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj01 t))) f0 f
- -> interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj02 t))) f0 g
- -> interp_type_rel_pointwise_uncurried_proj_from_option2 (t:=t) f0 f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj_from_option2, Morphisms.respectful_hetero.
- intros H0 H1 x.
- specialize (H0 (LiftOption.to' (Some x)) (SmartVarfMap proj01 x)).
- specialize (H1 (LiftOption.to' (Some x)) (SmartVarfMap proj02 x)).
- subst R.
- t proj012.
- Qed.
- End proj_from_option2.
- Global Arguments uncurry_interp_type_rel_pointwise_proj_from_option2 {_ _ _ _ _} proj {t f0 f g} _ _ _ _ _.
-
- Section proj1_from_option2.
- Context {interp_base_type0 interp_base_type1 : Type} {interp_base_type2 : base_type -> Type}
- (proj01 : interp_base_type0 -> interp_base_type1)
- (proj02 : forall t, interp_base_type0 -> interp_base_type2 t)
- (R : forall t, interp_base_type1 -> interp_base_type2 t -> Prop).
-
- Definition interp_type_rel_pointwise_uncurried_proj1_from_option2
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type0) t -> interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type interp_base_type2 t -> Prop
- := fun f0 f g
- => forall x : interp_flat_type (fun _ => interp_base_type0) (domain t),
- let x' := SmartVarfMap (fun _ => proj01) x in
- let y' := SmartVarfMap proj02 x in
- let fx := LiftOption.of' (f (LiftOption.to' (Some x'))) in
- let gy := g y' in
- let f0x := LiftOption.of' (f0 (LiftOption.to' (Some x))) in
- match f0x with
- | Some _ => True
- | None => False
- end
- -> match fx with
- | Some fx' => interp_flat_type_rel_pointwise (@R) fx' gy
- | None => True
- end.
-
- Lemma uncurry_interp_type_rel_pointwise_proj1_from_option2
- {t : type base_type}
- {f0}
- {f : interp_type (LiftOption.interp_base_type' interp_base_type1) t}
- {g : interp_type interp_base_type2 t}
- (proj012R : forall t, Reflexive (fun x y => @R _ (proj01 x) (proj02 t y)))
- : interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation2 (proj_eq_rel proj01)) f0 f
- -> interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj02 t))) f0 g
- -> interp_type_rel_pointwise_uncurried_proj1_from_option2 (t:=t) f0 f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj1_from_option2, Morphisms.respectful_hetero.
- intros H0 H1 x.
- specialize (H0 (LiftOption.to' (Some x)) (LiftOption.to' (Some (SmartVarfMap (fun _ => proj01) x)))).
- specialize (H1 (LiftOption.to' (Some x)) (SmartVarfMap proj02 x)).
- t proj012.
- Qed.
- End proj1_from_option2.
- Global Arguments uncurry_interp_type_rel_pointwise_proj1_from_option2 {_ _ _ _ _} R {t f0 f g} _ _ _ _ _.
-
- Section combine_related.
- Lemma related_flat_transitivity {interp_base_type1 interp_base_type2 interp_base_type3}
- {R1 : forall t : base_type, interp_base_type1 t -> interp_base_type2 t -> Prop}
- {R2 : forall t : base_type, interp_base_type1 t -> interp_base_type3 t -> Prop}
- {R3 : forall t : base_type, interp_base_type2 t -> interp_base_type3 t -> Prop}
- {t x y z}
- : (forall t a b c, (R1 t a b : Prop) -> (R2 t a c : Prop) -> (R3 t b c : Prop))
- -> interp_flat_type_rel_pointwise (t:=t) R1 x y
- -> interp_flat_type_rel_pointwise (t:=t) R2 x z
- -> interp_flat_type_rel_pointwise (t:=t) R3 y z.
- Proof.
- intro HRel; induction t; simpl; intuition eauto.
- Qed.
- End combine_related.
-End Relations.
diff --git a/src/Reflection/Z/Interpretations64.v b/src/Reflection/Z/Interpretations64.v
deleted file mode 100644
index 712cc96e0..000000000
--- a/src/Reflection/Z/Interpretations64.v
+++ /dev/null
@@ -1,13 +0,0 @@
-(** * Interpretation of PHOAS syntax for expression trees on ℤ *)
-Require Import Coq.ZArith.ZArith.
-Require Export Crypto.Reflection.Z.InterpretationsGen.
-Export Reflection.Syntax.Notations.
-
-Module BitSize64 <: BitSize.
- Definition bit_width : nat := 64%nat.
- Lemma bit_width_pos : (0 < Z.of_nat bit_width)%Z.
- Proof. simpl; omega. Qed.
-End BitSize64.
-
-Module Interpretations64 := InterpretationsGen BitSize64.
-Include Interpretations64.
diff --git a/src/Reflection/Z/Interpretations64/Relations.v b/src/Reflection/Z/Interpretations64/Relations.v
deleted file mode 100644
index 0a00ec10e..000000000
--- a/src/Reflection/Z/Interpretations64/Relations.v
+++ /dev/null
@@ -1,572 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.micromega.Psatz.
-Require Import Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Tuple.
-Require Import Crypto.Reflection.Z.InterpretationsGen.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.ModularArithmetic.ModularBaseSystemListZOperationsProofs.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.
-
-Definition proj_eq_rel {A B} (proj : A -> B) (x : A) (y : B) : Prop
- := proj x = y.
-Definition related'_Z {t} (x : BoundedWordW.BoundedWord) (y : Z.interp_base_type t) : Prop
- := proj_eq_rel (BoundedWordW.to_Z' _) x y.
-Definition related_Z t : BoundedWordW.interp_base_type t -> Z.interp_base_type t -> Prop
- := LiftOption.lift_relation (@related'_Z) t.
-Definition related'_wordW {t} (x : BoundedWordW.BoundedWord) (y : WordW.interp_base_type t) : Prop
- := proj_eq_rel (BoundedWordW.to_wordW' _) x y.
-Definition related_wordW t : BoundedWordW.interp_base_type t -> WordW.interp_base_type t -> Prop
- := LiftOption.lift_relation (@related'_wordW) t.
-Definition related_bounds t : BoundedWordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
- := LiftOption.lift_relation2 (proj_eq_rel BoundedWordW.BoundedWordToBounds) t.
-
-Definition related_wordW_Z t : WordW.interp_base_type t -> Z.interp_base_type t -> Prop
- := proj_eq_rel (WordW.to_Z _).
-
-Definition related'_wordW_bounds : WordW.wordW -> ZBounds.bounds -> Prop
- := fun value b => (0 <= Bounds.lower b /\ Bounds.lower b <= WordW.wordWToZ value <= Bounds.upper b /\ Z.log2 (Bounds.upper b) < Z.of_nat WordW.bit_width)%Z.
-Definition related_wordW_bounds : WordW.wordW -> ZBounds.t -> Prop
- := fun value b => match b with
- | Some b => related'_wordW_bounds value b
- | None => True
- end.
-Definition related_wordW_boundsi (t : base_type) : WordW.interp_base_type t -> ZBounds.interp_base_type t -> Prop
- := match t with
- | _ => related_wordW_bounds
- end.
-Definition related_wordW_boundsi' (t : base_type) : ZBounds.bounds -> WordW.interp_base_type t -> Prop
- := match t return ZBounds.bounds -> WordW.interp_base_type t -> Prop with
- | _ => fun x y => related'_wordW_bounds y x
- end.
-
-Local Notation related_op R interp_op1 interp_op2
- := (forall (src dst : flat_type base_type) (op : op src dst)
- (sv1 : interp_flat_type _ src) (sv2 : interp_flat_type _ src),
- interp_flat_type_rel_pointwise R sv1 sv2 ->
- interp_flat_type_rel_pointwise R (interp_op1 _ _ op sv1) (interp_op2 _ _ op sv2))
- (only parsing).
-Local Notation related_const R interp f g
- := (forall (t : base_type) (v : interp t), R t (f t v) (g t v))
- (only parsing).
-
-Local Ltac related_const_t :=
- let v := fresh in
- let t := fresh in
- intros t v; destruct t; intros; simpl in *; hnf; simpl;
- cbv [BoundedWordW.wordWToBoundedWord related'_Z LiftOption.of' related_Z related_wordW related'_wordW proj_eq_rel smart_interp_flat_map BoundedWordW.to_Z' WordW.to_Z BoundedWordW.to_wordW' BoundedWordW.of_wordW BoundedWordW.boundedWordToWordW] in *;
- break_innermost_match; simpl;
- first [ tauto
- | Z.ltb_to_lt;
- pose proof (WordW.wordWToZ_log_bound v);
- try omega ].
-
-Lemma related_Z_const : related_const related_Z WordW.interp_base_type BoundedWordW.of_wordW WordW.to_Z.
-Proof. related_const_t. Qed.
-Lemma related_bounds_const : related_const related_bounds WordW.interp_base_type BoundedWordW.of_wordW ZBounds.of_wordW.
-Proof. related_const_t. Qed.
-Lemma related_wordW_const : related_const related_wordW WordW.interp_base_type BoundedWordW.of_wordW (fun _ x => x).
-Proof. related_const_t. Qed.
-
-Local Ltac related_wordW_op_t_step :=
- first [ exact I
- | reflexivity
- | progress intros
- | progress inversion_option
- | progress ZBounds.inversion_bounds
- | progress subst
- | progress destruct_head' False
- | progress destruct_head' prod
- | progress destruct_head' and
- | progress destruct_head' option
- | progress destruct_head' BoundedWordW.BoundedWord
- | progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds'] in *
- | progress simpl @fst in *
- | progress simpl @snd in *
- | progress simpl @BoundedWord.upper in *
- | progress simpl @BoundedWord.lower in *
- | progress break_match
- | progress break_match_hyps
- | congruence
- | match goal with
- | [ H : ?op _ = Some _ |- _ ]
- => let H' := fresh in
- rename H into H';
- first [ pose proof (@BoundedWordW.t_map1_correct _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map2_correct _ _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map4_correct _ _ _ _ _ _ H') as H; clear H' ];
- simpl in H
- | [ H : ?op _ = None |- _ ]
- => let H' := fresh in
- rename H into H';
- first [ pose proof (@BoundedWordW.t_map1_correct_None _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map2_correct_None _ _ _ _ _ H') as H; clear H'
- | pose proof (@BoundedWordW.t_map4_correct_None _ _ _ _ _ H') as H; clear H' ];
- simpl in H
- end
- | progress cbv [related'_wordW proj_eq_rel BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value] in *
- | match goal with
- | [ H : ?op None _ = Some _ |- _ ] => progress simpl in H
- | [ H : ?op _ None = Some _ |- _ ] => progress simpl in H
- | [ H : ?op (Some _) (Some _) = Some _ |- _ ] => progress simpl in H
- | [ H : ?op (Some _) (Some _) = None |- _ ] => progress simpl in H
- end ].
-Local Ltac related_wordW_op_t := repeat related_wordW_op_t_step.
-
-Lemma related_wordW_t_map1 T opW opB pf
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Tbase T) related_wordW sv1 sv2
- -> @related_wordW T (BoundedWordW.t_map1 opW opB pf sv1) (opW sv2).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_wordW_t_map2 T opW opB pf
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Tbase T) (Tbase T)) related_wordW sv1 sv2
- -> @related_wordW T (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opW (fst sv2) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_wordW_t_map4 T opW opB pf
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Prod (Prod (Tbase T) (Tbase T)) (Tbase T)) (Tbase T)) related_wordW sv1 sv2
- -> @related_wordW T (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
- (opW (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_tuples_None_left
- T' n T interp_base_type'
- (R : forall t, LiftOption.interp_base_type' T t -> interp_base_type' t -> Prop)
- (RNone : forall v, R T' None v)
- (v : interp_flat_type interp_base_type' (tuple (Tbase T') (S n)))
- : interp_flat_type_rel_pointwise
- R
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option None))
- v.
-Proof.
- induction n; simpl; intuition.
-Qed.
-
-Lemma related_tuples_Some_left
- T' n T interp_base_type'
- (R : forall t, T -> interp_base_type' t -> Prop)
- u
- (v : interp_flat_type interp_base_type' (tuple (Tbase T') n))
- : interp_flat_type_rel_pointwise
- R
- (flat_interp_untuple (T:=Tbase T') u)
- v
- <-> interp_flat_type_rel_pointwise
- (LiftOption.lift_relation R)
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option (Some u)))
- v.
-Proof.
- destruct n as [|n]; [ reflexivity | ].
- induction n; [ reflexivity | ].
- simpl in *; rewrite <- IHn; clear IHn.
- reflexivity.
-Qed.
-
-Lemma related_tuples_Some_left_ext
- {T' : base_type} {n T interp_base_type'}
- {R : forall t, T -> interp_base_type' t -> Prop}
- {u v u'}
- (H : Tuple.lift_option (flat_interp_tuple (T:=Tbase T') (n:=n) u) = Some u')
- : interp_flat_type_rel_pointwise
- R
- (flat_interp_untuple (T:=Tbase T') u') v
- <-> interp_flat_type_rel_pointwise
- (LiftOption.lift_relation R)
- u v.
-Proof.
- destruct n as [|n]; [ reflexivity | ].
- induction n.
- { simpl in *; subst; reflexivity. }
- { destruct_head_hnf' prod.
- simpl in H; break_match_hyps; inversion_option; inversion_prod; subst.
- simpl; rewrite <- IHn by eassumption; clear IHn.
- reflexivity. }
-Qed.
-
-Lemma related_tuples_proj_eq_rel_untuple
- {T' : base_type} {n T interp_base_type'}
- {proj : forall t, T -> interp_base_type' t}
- {u : Tuple.tuple _ n} {v : Tuple.tuple _ n}
- : interp_flat_type_rel_pointwise
- (fun t => proj_eq_rel (proj t))
- (flat_interp_untuple (T:=Tbase T') u)
- (flat_interp_untuple (T:=Tbase T') v)
- <-> (Tuple.map (proj _) u = v).
-Proof.
- destruct n as [|n]; [ destruct_head_hnf' unit; simpl; tauto | ].
- induction n; [ reflexivity | ].
- destruct_head_hnf' prod.
- simpl @Tuple.tuple.
- rewrite !Tuple.map_S, path_prod_uncurried_iff, <- prod_iff_and; unfold fst, snd.
- rewrite <- IHn.
- reflexivity.
-Qed.
-
-Lemma related_tuples_proj_eq_rel_tuple
- {T' : base_type} {n T interp_base_type'}
- {proj : forall t, T -> interp_base_type' t}
- {u v}
- : interp_flat_type_rel_pointwise
- (fun t => proj_eq_rel (proj t))
- u v
- <-> (Tuple.map (proj _) (flat_interp_tuple (n:=n) (T:=Tbase T') u)
- = flat_interp_tuple (T:=Tbase T') v).
-Proof.
- rewrite <- related_tuples_proj_eq_rel_untuple, !flat_interp_untuple_tuple; reflexivity.
-Qed.
-
-Local Arguments LiftOption.lift_relation2 _ _ _ _ !_ !_ / .
-Lemma related_tuples_lift_relation2_untuple
- (T' : base_type) n T U
- (R : T -> U -> Prop)
- (t : option (Tuple.tuple T (S n)))
- (u : option (Tuple.tuple U (S n)))
- : interp_flat_type_rel_pointwise
- (LiftOption.lift_relation2 R)
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option t))
- (flat_interp_untuple (T:=Tbase T') (Tuple.push_option u))
- <-> LiftOption.lift_relation2
- (interp_flat_type_rel_pointwise (fun _ => R))
- T'
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => T) (T:=Tbase T')) t)
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => U) (T:=Tbase T')) u).
-Proof.
- induction n.
- { destruct_head' option; reflexivity. }
- { specialize (IHn (option_map (@fst _ _) t) (option_map (@fst _ _) u)).
- destruct_head' option;
- destruct_head_hnf' prod;
- simpl @option_map in *;
- simpl @LiftOption.lift_relation2 in *;
- try (rewrite <- IHn; reflexivity);
- try (simpl @interp_flat_type_rel_pointwise; tauto). }
-Qed.
-
-Lemma related_tuples_lift_relation2_untuple_ext
- (T' : base_type) {n T U}
- {R : T -> U -> Prop}
- {t u}
- (H : (exists v, Tuple.lift_option (n:=S n) (flat_interp_tuple (T:=Tbase T') t) = Some v)
- \/ (exists v, Tuple.lift_option (n:=S n) (flat_interp_tuple (T:=Tbase T') u) = Some v))
- : interp_flat_type_rel_pointwise
- (LiftOption.lift_relation2 R)
- t u
- <-> LiftOption.lift_relation2
- (interp_flat_type_rel_pointwise (fun _ => R))
- T'
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => T) (T:=Tbase T')) (Tuple.lift_option (flat_interp_tuple (T:=Tbase T') t)))
- (option_map (flat_interp_untuple (interp_base_type:=fun _ => U) (T:=Tbase T')) (Tuple.lift_option (flat_interp_tuple (T:=Tbase T') u))).
-Proof.
- induction n.
- { destruct_head_hnf' option; reflexivity. }
- { specialize (IHn (fst t) (fst u)).
- lazymatch type of IHn with
- | ?T -> _ => let H := fresh in assert (H : T); [ | specialize (IHn H); clear H ]
- end.
- { destruct_head' or; [ left | right ]; destruct_head' ex; destruct_head_hnf' prod; eexists;
- (etransitivity;
- [ | first [ refine (f_equal (option_map (@fst _ _)) (_ : _ = Some (_, _))); eassumption
- | refine (f_equal (option_map (@snd _ _)) (_ : _ = Some (_, _))); eassumption ] ]);
- instantiate; simpl in *; break_match; simpl in *; congruence. }
- destruct_head_hnf' prod;
- destruct_head_hnf' option;
- simpl @fst in *; simpl @snd in *;
- (etransitivity; [ simpl @interp_flat_type_rel_pointwise | reflexivity ]);
- try solve [ repeat first [ progress simpl in *
- | tauto
- | congruence
- | progress destruct_head ex
- | progress destruct_head or
- | progress break_match ] ]. }
-Qed.
-
-Lemma lift_option_flat_interp_tuple
- (T' : base_type) {n T x y}
- : (Tuple.lift_option (n:=S n) (A:=T) (flat_interp_tuple (interp_base_type:=LiftOption.interp_base_type' _) (T:=Tbase T') x) = Some y)
- <-> (x = flat_interp_untuple (T:=Tbase T') (n:=S n) (Tuple.push_option (n:=S n) (Some y))).
-Proof.
- rewrite Tuple.push_lift_option; generalize (Tuple.push_option (Some y)); intro.
- split; intro; subst;
- rewrite ?flat_interp_tuple_untuple, ?flat_interp_untuple_tuple;
- reflexivity.
-Qed.
-
-Lemma lift_option_None_interp_flat_type_rel_pointwise_1
- (T' : base_type) T U n R x y
- (H : interp_flat_type_rel_pointwise (LiftOption.lift_relation2 R) x y)
- (HNone : Tuple.lift_option (A:=T) (n:=S n) (flat_interp_tuple (T:=Tbase T') (n:=S n) x) = None)
- : Tuple.lift_option (A:=U) (n:=S n) (flat_interp_tuple (T:=Tbase T') (n:=S n) y) = None.
-Proof.
- induction n; [ | specialize (IHn (fst x) (fst y) (proj1 H)) ];
- repeat first [ progress destruct_head_hnf' False
- | reflexivity
- | progress inversion_option
- | progress simpl in *
- | progress subst
- | progress specialize_by congruence
- | progress destruct_head_hnf' prod
- | progress destruct_head_hnf' and
- | progress destruct_head_hnf' option
- | progress break_match
- | progress break_match_hyps ].
-Qed.
-
-Local Arguments LiftOption.lift_relation _ _ _ _ !_ _ / .
-Local Arguments LiftOption.of' _ _ !_ / .
-Local Arguments BoundedWordW.BoundedWordToBounds !_ / .
-
-Local Ltac unfold_related_const :=
- intros; hnf; simpl;
- cbv [related_wordW LiftOption.lift_relation LiftOption.of' BoundedWordW.wordWToBoundedWord BoundedWordW.SmartBuildBoundedWord BoundedWordW.of_Z BoundedWordW.of_wordW BoundedWordW.wordWToBoundedWord BoundedWordW.to_Z' BoundedWordW.to_wordW'];
- simpl.
-
-Lemma related_wordW_op : related_op related_wordW (@BoundedWordW.interp_op) (@WordW.interp_op).
-Proof.
- (let op := fresh in intros ?? op; destruct op; simpl);
- try first [ apply related_wordW_t_map1
- | apply related_wordW_t_map2
- | apply related_wordW_t_map4
- | unfold_related_const; break_innermost_match; reflexivity
- | exact (fun _ _ x => x) ].
-Qed.
-
-Lemma related_bounds_t_map1 T opW opB pf
- (HN : opB None = None)
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Tbase T) related_bounds sv1 sv2
- -> @related_bounds T (BoundedWordW.t_map1 opW opB pf sv1) (opB sv2).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_bounds_t_map2 T opW opB pf
- (HN0 : forall v, opB None v = None)
- (HN1 : forall v, opB v None = None)
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Tbase T) (Tbase T)) related_bounds sv1 sv2
- -> @related_bounds T (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opB (fst sv2) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_wordW_op_t.
-Qed.
-
-Lemma related_bounds_t_map4 T opW opB pf
- (HN0 : forall x y z, opB None x y z = None)
- (HN1 : forall x y z, opB x None y z = None)
- (HN2 : forall x y z, opB x y None z = None)
- (HN3 : forall x y z, opB x y z None = None)
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Prod (Prod (Tbase T) (Tbase T)) (Tbase T)) (Tbase T)) related_bounds sv1 sv2
- -> @related_bounds T (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
- (opB (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- destruct_head prod.
- intros; destruct_head' prod.
- progress cbv [related_wordW related_bounds related_Z LiftOption.lift_relation LiftOption.lift_relation2 LiftOption.of' smart_interp_flat_map BoundedWordW.BoundedWordToBounds BoundedWordW.to_bounds' proj_eq_rel] in *.
- destruct_head' option; destruct_head_hnf' and; destruct_head_hnf' False; subst;
- try solve [ simpl; rewrite ?HN0, ?HN1, ?HN2, ?HN3; tauto ];
- [].
- related_wordW_op_t.
-Qed.
-
-Local Arguments Tuple.lift_option : simpl never.
-Local Arguments Tuple.push_option : simpl never.
-Local Arguments Tuple.map : simpl never.
-Local Arguments Tuple.map2 : simpl never.
-
-Local Arguments ZBounds.SmartBuildBounds _ _ / .
-
-Local Ltac related_const_op_t :=
- unfold_related_const; break_innermost_match; try reflexivity; hnf; simpl;
- repeat match goal with
- | [ H : andb _ _ = true |- _ ] => apply andb_prop in H
- | _ => progress destruct_head' and
- | _ => progress Z.ltb_to_lt
- | _ => rewrite WordW.wordWToZ_ZToWordW by (simpl @Z.of_nat; omega)
- | [ H : _ |- _ ] => rewrite WordW.wordWToZ_ZToWordW in H by (simpl @Z.of_nat; omega)
- | [ H : (Z.log2 ?x < ?y)%Z |- _ ]
- => unique assert (x < 2^y)%Z by (apply WordW.log2_lt_pow2_alt_proj_r2l; omega)
- | _ => reflexivity
- | _ => omega
- end.
-
-Lemma related_bounds_op : related_op related_bounds (@BoundedWordW.interp_op) (@ZBounds.interp_op).
-Proof.
- let op := fresh in intros ?? op; destruct op; simpl.
- { related_const_op_t. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map2; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map1; intros; destruct_head' option; unfold ZBounds.neg; break_match; reflexivity. }
- { apply related_bounds_t_map4; intros; destruct_head' option; reflexivity. }
- { apply related_bounds_t_map4; intros; destruct_head' option; reflexivity. }
- { exact (fun _ _ x => x). }
-Qed.
-
-(* we [match] and [eexact] rather than [eassumption] so that we can backtrack across hypothesis choice, so that we're hypothesis-order-independent *)
-Local Ltac WordW.Rewrites.wordW_util_arith ::=
- solve [ autorewrite with Zshift_to_pow; omega
- | autorewrite with Zshift_to_pow; nia
- | autorewrite with Zshift_to_pow; auto with zarith
- | match goal with
- | [ H : _ |- _ ]
- => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono | eexact H ];
- autorewrite with Zshift_to_pow; auto using Z.mul_le_mono_nonneg with zarith;
- solve [ omega
- | nia
- | etransitivity; [ eapply Z.div_le_mono | eapply Z.div_le_compat_l ];
- auto with zarith ]
- end
- | apply Z.land_nonneg; WordW.Rewrites.wordW_util_arith
- | match goal with
- | [ H : _ |- _ ]
- => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono | eexact H ];
- instantiate; apply Z.min_case_strong; intros;
- first [ etransitivity; [ apply Z.land_upper_bound_l | ]; omega
- | etransitivity; [ apply Z.land_upper_bound_r | ]; omega ]
- end
- | rewrite Z.log2_lor by omega;
- apply Z.max_case_strong; intro;
- match goal with
- | [ H : _ |- _ ]
- => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono; eexact H | assumption ]
- end
- | eapply Z.le_lt_trans; [ eapply Z.log2_le_mono, neg_upperbound | ];
- WordW.Rewrites.wordW_util_arith
- | (progress unfold ModularBaseSystemListZOperations.cmovne, ModularBaseSystemListZOperations.cmovl, ModularBaseSystemListZOperations.neg); break_match;
- WordW.Rewrites.wordW_util_arith ].
-Local Ltac related_Z_op_t_step :=
- first [ progress related_wordW_op_t_step
- | progress cbv [related'_Z proj_eq_rel BoundedWordW.to_Z' BoundedWordW.to_wordW' WordW.to_Z BoundedWordW.boundedWordToWordW BoundedWord.value] in *
- | autorewrite with push_wordWToZ ].
-Local Ltac related_Z_op_t := repeat related_Z_op_t_step.
-
-Local Notation is_bounded_by value lower upper
- := ((0 <= lower /\ lower <= WordW.wordWToZ value <= upper /\ Z.log2 upper < Z.of_nat WordW.bit_width)%Z)
- (only parsing).
-Local Notation is_in_bounds value bounds
- := (is_bounded_by value (Bounds.lower bounds) (Bounds.upper bounds))
- (only parsing).
-
-Lemma related_Z_t_map1 T opZ opW opB pf
- (H : forall x bxs brs,
- Some brs = opB (Some bxs)
- -> is_in_bounds x bxs
- -> is_in_bounds (opW x) brs
- -> WordW.wordWToZ (opW x) = (opZ (WordW.wordWToZ x)))
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Tbase T) related_Z sv1 sv2
- -> @related_Z T (BoundedWordW.t_map1 opW opB pf sv1) (opZ sv2).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_Z_op_t.
- eapply H; eauto.
-Qed.
-
-Lemma related_Z_t_map2 T opZ opW opB pf
- (H : forall x y bxs bys brs,
- Some brs = opB (Some bxs) (Some bys)
- -> is_in_bounds x bxs
- -> is_in_bounds y bys
- -> is_in_bounds (opW x y) brs
- -> WordW.wordWToZ (opW x y) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y)))
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=Prod (Tbase T) (Tbase T)) related_Z sv1 sv2
- -> @related_Z T (BoundedWordW.t_map2 opW opB pf (fst sv1) (snd sv1)) (opZ (fst sv2) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_Z_op_t.
- eapply H; eauto.
-Qed.
-
-Lemma related_Z_t_map4 T opZ opW opB pf
- (H : forall x y z w bxs bys bzs bws brs,
- Some brs = opB (Some bxs) (Some bys) (Some bzs) (Some bws)
- -> is_in_bounds x bxs
- -> is_in_bounds y bys
- -> is_in_bounds z bzs
- -> is_in_bounds w bws
- -> is_in_bounds (opW x y z w) brs
- -> WordW.wordWToZ (opW x y z w) = (opZ (WordW.wordWToZ x) (WordW.wordWToZ y) (WordW.wordWToZ z) (WordW.wordWToZ w)))
- sv1 sv2
- : interp_flat_type_rel_pointwise (t:=(Tbase T * Tbase T * Tbase T * Tbase T)%ctype) related_Z sv1 sv2
- -> @related_Z T (BoundedWordW.t_map4 opW opB pf (fst (fst (fst sv1))) (snd (fst (fst sv1))) (snd (fst sv1)) (snd sv1))
- (opZ (fst (fst (fst sv2))) (snd (fst (fst sv2))) (snd (fst sv2)) (snd sv2)).
-Proof.
- cbv [interp_flat_type BoundedWordW.interp_base_type ZBounds.interp_base_type LiftOption.interp_base_type' interp_flat_type_rel_pointwise interp_flat_type_rel_pointwise_gen_Prop] in *.
- related_Z_op_t.
- eapply H; eauto.
-Qed.
-
-Local Arguments related_Z _ !_ _ / .
-
-Local Arguments related'_Z _ _ _ / .
-
-Local Ltac related_Z_op_fin_t_step :=
- first [ progress subst
- | progress inversion_option
- | progress ZBounds.inversion_bounds
- | progress destruct_head' Bounds.bounds
- | progress destruct_head' ZBounds.bounds
- | progress destruct_head' and
- | progress simpl in * |-
- | progress break_match_hyps
- | congruence
- | progress inversion_option
- | intro
- | progress autorewrite with push_wordWToZ
- | match goal with
- | [ H : andb _ _ = true |- _ ] => rewrite Bool.andb_true_iff in H
- | [ H : context[Tuple.lift_option (Tuple.push_option _)] |- _ ]
- => rewrite Tuple.lift_push_option in H
- end
- | progress Z.ltb_to_lt ].
-Local Ltac related_Z_op_fin_t := repeat related_Z_op_fin_t_step.
-
-Local Opaque WordW.bit_width.
-
-Local Arguments ZBounds.neg _ !_ / .
-
-Lemma related_Z_op : related_op related_Z (@BoundedWordW.interp_op) (@Z.interp_op).
-Proof.
- let op := fresh in intros ?? op; destruct op; simpl.
- { related_const_op_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map2; related_Z_op_fin_t. }
- { apply related_Z_t_map1; related_Z_op_fin_t. }
- { apply related_Z_t_map4; related_Z_op_fin_t. }
- { apply related_Z_t_map4; related_Z_op_fin_t. }
- { exact (fun _ _ x => x). }
-Qed.
-
-Create HintDb interp_related discriminated.
-Hint Resolve related_Z_op related_bounds_op related_wordW_op related_Z_const related_bounds_const related_wordW_const : interp_related.
diff --git a/src/Reflection/Z/Interpretations64/RelationsCombinations.v b/src/Reflection/Z/Interpretations64/RelationsCombinations.v
deleted file mode 100644
index adfceb899..000000000
--- a/src/Reflection/Z/Interpretations64/RelationsCombinations.v
+++ /dev/null
@@ -1,253 +0,0 @@
-Require Import Coq.ZArith.ZArith.
-Require Import Coq.Classes.RelationClasses.
-Require Import Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Z.InterpretationsGen.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.Reflection.Z.Interpretations64.Relations.
-Require Import Crypto.Util.Option.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.
-
-Module Relations.
- Section proj.
- Context {interp_base_type1 interp_base_type2}
- (proj : forall t : base_type, interp_base_type1 t -> interp_base_type2 t).
-
- Let R {t : flat_type base_type} f g :=
- SmartVarfMap (t:=t) proj f = g.
-
- Definition interp_type_rel_pointwise_uncurried_proj
- {t : type base_type}
- : interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> Prop
- := fun f g
- => forall x : interp_flat_type interp_base_type1 (domain t),
- let y := SmartVarfMap proj x in
- let fx := f x in
- let gy := g y in
- @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj
- {t : type base_type}
- {f}
- {g}
- : (forall x y, @R (domain t) x y -> @R (codomain t) (f x) (g y))
- -> interp_type_rel_pointwise_uncurried_proj (t:=t) f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj.
- destruct t as [A B]; simpl in *.
- subst R; simpl in *.
- eauto.
- Qed.
- End proj.
-
- Section proj_option.
- Context {interp_base_type1 : Type} {interp_base_type2 : base_type -> Type}
- (proj_option : forall t, interp_base_type1 -> interp_base_type2 t).
-
- Let R {t : flat_type base_type} f g :=
- let f' := LiftOption.of' (t:=t) f in
- match f' with
- | Some f' => SmartVarfMap proj_option f' = g
- | None => True
- end.
-
- Definition interp_type_rel_pointwise_uncurried_proj_option
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type interp_base_type2 t -> Prop
- := fun f g
- => forall x : interp_flat_type (fun _ => interp_base_type1) (domain t),
- let y := SmartVarfMap proj_option x in
- let fx := f (LiftOption.to' (Some x)) in
- let gy := g y in
- @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj_option
- {t : type base_type}
- {f}
- {g}
- : (forall x y, @R (domain t) x y -> @R (codomain t) (f x) (g y))
- -> interp_type_rel_pointwise_uncurried_proj_option (t:=t) f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj_option.
- destruct t as [A B]; simpl in *.
- subst R; simpl in *.
- intros H x; apply H; simpl.
- rewrite LiftOption.of'_to'; reflexivity.
- Qed.
- End proj_option.
-
- Section proj_option2.
- Context {interp_base_type1 : Type} {interp_base_type2 : Type}
- (proj : interp_base_type1 -> interp_base_type2).
-
- Let R {t : flat_type base_type} f g :=
- let f' := LiftOption.of' (t:=t) f in
- let g' := LiftOption.of' (t:=t) g in
- match f', g' with
- | Some f', Some g' => SmartVarfMap (fun _ => proj) f' = g'
- | None, None => True
- | Some _, _ => False
- | None, _ => False
- end.
-
- Definition interp_type_rel_pointwise_uncurried_proj_option2
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type (LiftOption.interp_base_type' interp_base_type2) t -> Prop
- := fun f g
- => forall x : interp_flat_type (fun _ => interp_base_type1) (domain t),
- let y := SmartVarfMap (fun _ => proj) x in
- let fx := f (LiftOption.to' (Some x)) in
- let gy := g (LiftOption.to' (Some y)) in
- @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj_option2
- {t : type base_type}
- {f}
- {g}
- : (forall x y, @R (domain t) x y -> @R (codomain t) (f x) (g y))
- -> interp_type_rel_pointwise_uncurried_proj_option2 (t:=t) f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj_option2.
- destruct t as [A B]; simpl in *.
- subst R; simpl in *.
- intros H x; apply H; simpl.
- rewrite !LiftOption.of'_to'; reflexivity.
- Qed.
- End proj_option2.
-
- Local Ltac t proj012 :=
- repeat match goal with
- | _ => progress cbv beta in *
- | _ => progress break_match; try tauto; []
- | _ => progress subst
- | [ H : _ |- _ ]
- => first [ rewrite !LiftOption.lift_relation_flat_type_pointwise in H
- by (eassumption || rewrite LiftOption.of'_to'; reflexivity)
- | rewrite !LiftOption.lift_relation2_flat_type_pointwise in H
- by (eassumption || rewrite LiftOption.of'_to'; reflexivity)
- | rewrite !@lift_interp_flat_type_rel_pointwise_f_eq_id2 in H
- | rewrite !@lift_interp_flat_type_rel_pointwise_f_eq in H ]
- | _ => progress unfold proj_eq_rel in *
- | _ => progress specialize_by reflexivity
- | _ => rewrite SmartVarfMap_compose
- | _ => setoid_rewrite proj012
- | [ |- context G[fun x y => ?f x y] ]
- => let G' := context G[f] in change G'
- | [ |- context G[fun (_ : ?T) x => ?f x] ]
- => let G' := context G[fun _ : T => f] in change G'
- | [ H : context G[fun (_ : ?T) x => ?f x] |- _ ]
- => let G' := context G[fun _ : T => f] in change G' in H
- | _ => rewrite_hyp <- !*; []
- | _ => reflexivity
- | _ => rewrite interp_flat_type_rel_pointwise_SmartVarfMap
- end.
-
- Section proj_from_option2.
- Context {interp_base_type0 : Type} {interp_base_type1 interp_base_type2 : base_type -> Type}
- (proj01 : forall t, interp_base_type0 -> interp_base_type1 t)
- (proj02 : forall t, interp_base_type0 -> interp_base_type2 t)
- (proj : forall t, interp_base_type1 t -> interp_base_type2 t).
-
- Let R {t : flat_type base_type} f g :=
- proj_eq_rel (SmartVarfMap proj (t:=t)) f g.
-
- Definition interp_type_rel_pointwise_uncurried_proj_from_option2
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type0) t -> interp_type interp_base_type1 t -> interp_type interp_base_type2 t -> Prop
- := fun f0 f g
- => forall x : interp_flat_type (fun _ => interp_base_type0) (domain t),
- let x' := SmartVarfMap proj01 x in
- let y' := SmartVarfMap proj x' in
- let fx := f x' in
- let gy := g y' in
- let f0x := LiftOption.of' (f0 (LiftOption.to' (Some x))) in
- match f0x with
- | Some _ => True
- | None => False
- end
- -> @R _ fx gy.
-
- Lemma uncurry_interp_type_rel_pointwise_proj_from_option2
- {t : type base_type}
- {f0}
- {f : interp_type interp_base_type1 t}
- {g : interp_type interp_base_type2 t}
- (proj012 : forall t x, proj t (proj01 t x) = proj02 t x)
- : interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj01 t))) f0 f
- -> interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj02 t))) f0 g
- -> interp_type_rel_pointwise_uncurried_proj_from_option2 (t:=t) f0 f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj_from_option2, Morphisms.respectful_hetero.
- intros H0 H1 x.
- specialize (H0 (LiftOption.to' (Some x)) (SmartVarfMap proj01 x)).
- specialize (H1 (LiftOption.to' (Some x)) (SmartVarfMap proj02 x)).
- subst R.
- t proj012.
- Qed.
- End proj_from_option2.
- Global Arguments uncurry_interp_type_rel_pointwise_proj_from_option2 {_ _ _ _ _} proj {t f0 f g} _ _ _ _ _.
-
- Section proj1_from_option2.
- Context {interp_base_type0 interp_base_type1 : Type} {interp_base_type2 : base_type -> Type}
- (proj01 : interp_base_type0 -> interp_base_type1)
- (proj02 : forall t, interp_base_type0 -> interp_base_type2 t)
- (R : forall t, interp_base_type1 -> interp_base_type2 t -> Prop).
-
- Definition interp_type_rel_pointwise_uncurried_proj1_from_option2
- {t : type base_type}
- : interp_type (LiftOption.interp_base_type' interp_base_type0) t -> interp_type (LiftOption.interp_base_type' interp_base_type1) t -> interp_type interp_base_type2 t -> Prop
- := fun f0 f g
- => forall x : interp_flat_type (fun _ => interp_base_type0) (domain t),
- let x' := SmartVarfMap (fun _ => proj01) x in
- let y' := SmartVarfMap proj02 x in
- let fx := LiftOption.of' (f (LiftOption.to' (Some x'))) in
- let gy := g y' in
- let f0x := LiftOption.of' (f0 (LiftOption.to' (Some x))) in
- match f0x with
- | Some _ => True
- | None => False
- end
- -> match fx with
- | Some fx' => interp_flat_type_rel_pointwise (@R) fx' gy
- | None => True
- end.
-
- Lemma uncurry_interp_type_rel_pointwise_proj1_from_option2
- {t : type base_type}
- {f0}
- {f : interp_type (LiftOption.interp_base_type' interp_base_type1) t}
- {g : interp_type interp_base_type2 t}
- (proj012R : forall t, Reflexive (fun x y => @R _ (proj01 x) (proj02 t y)))
- : interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation2 (proj_eq_rel proj01)) f0 f
- -> interp_type_rel_pointwise (t:=t) (LiftOption.lift_relation (fun t => proj_eq_rel (proj02 t))) f0 g
- -> interp_type_rel_pointwise_uncurried_proj1_from_option2 (t:=t) f0 f g.
- Proof.
- unfold interp_type_rel_pointwise_uncurried_proj1_from_option2, Morphisms.respectful_hetero.
- intros H0 H1 x.
- specialize (H0 (LiftOption.to' (Some x)) (LiftOption.to' (Some (SmartVarfMap (fun _ => proj01) x)))).
- specialize (H1 (LiftOption.to' (Some x)) (SmartVarfMap proj02 x)).
- t proj012.
- Qed.
- End proj1_from_option2.
- Global Arguments uncurry_interp_type_rel_pointwise_proj1_from_option2 {_ _ _ _ _} R {t f0 f g} _ _ _ _ _.
-
- Section combine_related.
- Lemma related_flat_transitivity {interp_base_type1 interp_base_type2 interp_base_type3}
- {R1 : forall t : base_type, interp_base_type1 t -> interp_base_type2 t -> Prop}
- {R2 : forall t : base_type, interp_base_type1 t -> interp_base_type3 t -> Prop}
- {R3 : forall t : base_type, interp_base_type2 t -> interp_base_type3 t -> Prop}
- {t x y z}
- : (forall t a b c, (R1 t a b : Prop) -> (R2 t a c : Prop) -> (R3 t b c : Prop))
- -> interp_flat_type_rel_pointwise (t:=t) R1 x y
- -> interp_flat_type_rel_pointwise (t:=t) R2 x z
- -> interp_flat_type_rel_pointwise (t:=t) R3 y z.
- Proof.
- intro HRel; induction t; simpl; intuition eauto.
- Qed.
- End combine_related.
-End Relations.
diff --git a/src/Reflection/Z/InterpretationsGen.v b/src/Reflection/Z/InterpretationsGen.v
deleted file mode 100644
index 68e71e022..000000000
--- a/src/Reflection/Z/InterpretationsGen.v
+++ /dev/null
@@ -1,915 +0,0 @@
-(** * Interpretation of PHOAS syntax for expression trees on ℤ *)
-Require Import Coq.ZArith.ZArith.
-Require Import Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.ModularArithmetic.ModularBaseSystemListZOperations.
-Require Import Crypto.Util.Equality.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Option.
-Require Crypto.Util.Tuple.
-Require Crypto.Util.HList.
-Require Import Crypto.Util.Bool.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Tactics.
-Require Import Bedrock.Word.
-Require Import Crypto.Util.WordUtil.
-Export Reflection.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).
-
-Module Type BitSize.
- Parameter bit_width : nat.
- Axiom bit_width_pos : (0 < Z.of_nat bit_width)%Z.
-End BitSize.
-
-Module Import Bounds.
- Record bounds := { lower : Z ; upper : Z }.
-End Bounds.
-Module Import BoundedWord.
- Record BoundedWordGen wordW (is_bounded_by : _ -> _ -> _ -> Prop) :=
- { lower : Z ; value : wordW ; upper : Z ;
- in_bounds : is_bounded_by value lower upper }.
- Global Arguments lower {_ _} _.
- Global Arguments value {_ _} _.
- Global Arguments upper {_ _} _.
- Global Arguments in_bounds {_ _} _.
-End BoundedWord.
-Module InterpretationsGen (Bit : BitSize).
- Module Z.
- Definition interp_base_type (t : base_type) : Type := interp_base_type t.
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := interp_op src dst f.
- End Z.
-
- Module LiftOption.
- Section lift_option.
- Context (T : Type).
-
- Definition interp_flat_type (t : flat_type base_type)
- := option (interp_flat_type (fun _ => T) t).
-
- Definition interp_base_type' (t : base_type)
- := option T.
-
- Definition of' {t} : Syntax.interp_flat_type interp_base_type' t -> interp_flat_type t
- := @smart_interp_flat_map
- base_type
- interp_base_type' interp_flat_type
- (fun t => match t with _ => fun x => x end)
- (Some tt)
- (fun _ _ x y => match x, y with
- | Some x', Some y' => Some (x', y')
- | _, _ => None
- end)
- t.
-
- Lemma of'_pair {A B} x
- : @of' (A * B) x = match of' (fst x), of' (snd x) with
- | Some x', Some y' => Some (x', y')
- | _, _ => None
- end.
- Proof. reflexivity. Qed.
-
- Fixpoint to' {t} : interp_flat_type t -> Syntax.interp_flat_type interp_base_type' t
- := match t return interp_flat_type t -> Syntax.interp_flat_type interp_base_type' t with
- | Tbase _ => fun x => x
- | Unit => fun _ => tt
- | Prod A B => fun x => (@to' A (option_map (@fst _ _) x),
- @to' B (option_map (@snd _ _) x))
- end.
-
- Lemma of'_to' {t} v : of' (@to' t (Some v)) = Some v.
- Proof.
- induction t;
- repeat first [ progress simpl
- | progress destruct_head_hnf prod
- | progress destruct_head_hnf unit
- | progress destruct_head base_type
- | rewrite of'_pair
- | rewrite_hyp !*
- | reflexivity ].
- Qed.
-
- Section lift_relation.
- Context {interp_base_type2 : base_type -> Type}
- (R : forall t, T -> interp_base_type2 t -> Prop).
- Definition lift_relation
- : forall t, interp_base_type' t -> interp_base_type2 t -> Prop
- := fun t x y => match of' (t:=Tbase t) x with
- | Some x' => R t x' y
- | None => True
- end.
-
- Lemma lift_relation_flat_type_pointwise t x y x'
- (Hx : of' x = Some x')
- : interp_flat_type_rel_pointwise lift_relation x y
- <-> interp_flat_type_rel_pointwise (t:=t) R x' y.
- Proof.
- induction t; simpl; try tauto.
- { unfold lift_relation; rewrite Hx; reflexivity. }
- { rewrite of'_pair in Hx.
- destruct (of' (fst x)) eqn:?, (of' (snd x)) eqn:?; try congruence.
- inversion_option; subst.
- destruct_head_hnf prod; split_iff; intuition eauto. }
- Qed.
-
- Lemma lift_relation_type_pointwise t f g f'
- (Hx : forall x x', of' x = Some x' -> of' (f x) = Some (f' x'))
- : interp_type_rel_pointwise lift_relation f g
- -> interp_type_rel_pointwise (t:=t) R f' g.
- Proof.
- destruct t; simpl; unfold Morphisms.respectful_hetero.
- intros H x y p; specialize (H (to' (Some x)) y).
- eapply lift_relation_flat_type_pointwise in p; [ | apply of'_to'.. ].
- specialize (H p).
- eapply lift_relation_flat_type_pointwise in H; [ exact H | ].
- erewrite Hx; [ reflexivity | ].
- rewrite of'_to'; reflexivity.
- Qed.
- End lift_relation.
- End lift_option.
- Global Arguments of' {T t} _.
- Global Arguments to' {T t} _.
- Global Arguments lift_relation {T _} R _ _ _.
-
- Section lift_option2.
- Context (T U : Type) (R : T -> U -> Prop).
- Definition lift_relation2
- : forall t, interp_base_type' T t -> interp_base_type' U t -> Prop
- := fun t x y => match of' (t:=Tbase t) x, of' (t:=Tbase t) y with
- | Datatypes.Some x', Datatypes.Some y' => R x' y'
- | None, None => True
- | _, _ => False
- end.
-
- Lemma lift_relation2_flat_type_pointwise t x y x' y'
- (Hx : of' x = Datatypes.Some x')
- (Hy : of' y = Datatypes.Some y')
- : interp_flat_type_rel_pointwise lift_relation2 x y
- <-> interp_flat_type_rel_pointwise (t:=t) (fun _ => R) x' y'.
- Proof.
- induction t; simpl; try tauto.
- { unfold lift_relation2; rewrite Hx, Hy; reflexivity. }
- { rewrite of'_pair in Hx, Hy.
- destruct (of' (fst x)) eqn:?, (of' (snd x)) eqn:?; try congruence.
- destruct (of' (fst y)) eqn:?, (of' (snd y)) eqn:?; try congruence.
- inversion_option; subst.
- destruct_head_hnf prod; split_iff; intuition eauto. }
- Qed.
-
- Lemma lift_relation2_type_pointwise t f g f' g'
- (Hx : forall x x', of' x = Datatypes.Some x' -> of' (f x) = Datatypes.Some (f' x'))
- (Hy : forall x x', of' x = Datatypes.Some x' -> of' (g x) = Datatypes.Some (g' x'))
- : interp_type_rel_pointwise lift_relation2 f g
- -> interp_type_rel_pointwise (t:=t) (fun _ => R) f' g'.
- Proof.
- destruct t; simpl; unfold Morphisms.respectful_hetero.
- intros H x y p; specialize (H (to' (Datatypes.Some x)) (to' (Datatypes.Some y))).
- eapply lift_relation2_flat_type_pointwise in p; [ | apply of'_to'.. ].
- specialize (H p).
- eapply lift_relation2_flat_type_pointwise in H; [ exact H | .. ];
- [ erewrite Hx; [ reflexivity | ]
- | erewrite Hy; [ reflexivity | ] ];
- rewrite of'_to'; reflexivity.
- Qed.
- End lift_option2.
- Global Arguments lift_relation2 {T U} R _ _ _.
- End LiftOption.
-
- Module WordW.
- Include Bit.
- Definition wordW := word bit_width.
- Delimit Scope wordW_scope with wordW.
- Bind Scope wordW_scope with wordW.
-
- Definition wordWToZ (x : wordW) : Z
- := Z.of_N (wordToN x).
- Definition ZToWordW (x : Z) : wordW
- := NToWord _ (Z.to_N x).
-
- Ltac fold_WordW_Z :=
- repeat match goal with
- | [ |- context G[NToWord bit_width (Z.to_N ?x)] ]
- => let G' := context G [ZToWordW x] in change G'
- | [ |- context G[Z.of_N (wordToN ?x)] ]
- => let G' := context G [wordWToZ x] in change G'
- | [ H : context G[NToWord bit_width (Z.to_N ?x)] |- _ ]
- => let G' := context G [ZToWordW x] in change G' in H
- | [ H : context G[Z.of_N (wordToN ?x)] |- _ ]
- => let G' := context G [wordWToZ x] in change G' in H
- end.
-
- Create HintDb push_wordWToZ discriminated.
- Hint Extern 1 => progress autorewrite with push_wordWToZ in * : push_wordWToZ.
-
- Local Hint Resolve bit_width_pos : zarith.
-
- Ltac arith := solve [ omega | auto using bit_width_pos with zarith ].
-
- Lemma wordWToZ_bound w : (0 <= wordWToZ w < 2^Z.of_nat bit_width)%Z.
- Proof.
- pose proof (wordToNat_bound w) as H.
- apply Nat2Z.inj_lt in H.
- rewrite Zpow_pow2, Z2Nat.id in H by (apply Z.pow_nonneg; omega).
- unfold wordWToZ.
- rewrite wordToN_nat, nat_N_Z; omega.
- Qed.
-
- Lemma wordWToZ_log_bound w : (0 <= wordWToZ w /\ Z.log2 (wordWToZ w) < Z.of_nat bit_width)%Z.
- Proof.
- pose proof (wordWToZ_bound w) as H.
- destruct (Z_zerop (wordWToZ w)) as [H'|H'].
- { rewrite H'; simpl; split; auto with zarith. }
- { split; [ | apply Z.log2_lt_pow2 ]; omega. }
- Qed.
-
- Lemma ZToWordW_wordWToZ (x : wordW) : ZToWordW (wordWToZ x) = x.
- Proof.
- unfold ZToWordW, wordWToZ.
- rewrite N2Z.id, NToWord_wordToN.
- reflexivity.
- Qed.
- Hint Rewrite ZToWordW_wordWToZ : push_wordWToZ.
-
- Lemma wordWToZ_ZToWordW (x : Z) : (0 <= x < 2^Z.of_nat bit_width)%Z -> wordWToZ (ZToWordW x) = x.
- Proof.
- unfold ZToWordW, wordWToZ; intros [H0 H1].
- pose proof H1 as H1'; apply Z2Nat.inj_lt in H1'; [ | omega.. ].
- rewrite <- Z.pow_Z2N_Zpow in H1' by omega.
- replace (Z.to_nat 2) with 2%nat in H1' by reflexivity.
- rewrite wordToN_NToWord_idempotent, Z2N.id by (omega || auto using bound_check_nat_N).
- reflexivity.
- Qed.
- Hint Rewrite wordWToZ_ZToWordW using arith : push_wordWToZ.
-
- Definition add : wordW -> wordW -> wordW := @wplus _.
- Definition sub : wordW -> wordW -> wordW := @wminus _.
- Definition mul : wordW -> wordW -> wordW := @wmult _.
- Definition shl : wordW -> wordW -> wordW := @wordBin N.shiftl _.
- Definition shr : wordW -> wordW -> wordW := @wordBin N.shiftr _.
- Definition land : wordW -> wordW -> wordW := @wand _.
- Definition lor : wordW -> wordW -> wordW := @wor _.
- Definition neg (int_width : Z) : wordW -> wordW (* TODO: Is this right? *)
- := fun x => ZToWordW (ModularBaseSystemListZOperations.neg int_width (wordWToZ x)).
- Definition cmovne : wordW -> wordW -> wordW -> wordW -> wordW (* TODO: Is this right? *)
- := fun x y z w => ZToWordW (ModularBaseSystemListZOperations.cmovne (wordWToZ x) (wordWToZ y) (wordWToZ z) (wordWToZ w)).
- Definition cmovle : wordW -> wordW -> wordW -> wordW -> wordW (* TODO: Is this right? *)
- := fun x y z w => ZToWordW (ModularBaseSystemListZOperations.cmovl (wordWToZ x) (wordWToZ y) (wordWToZ z) (wordWToZ w)).
-
- Infix "+" := add : wordW_scope.
- Infix "-" := sub : wordW_scope.
- Infix "*" := mul : wordW_scope.
- Infix "<<" := shl : wordW_scope.
- Infix ">>" := shr : wordW_scope.
- Infix "&'" := land : wordW_scope.
-
- (*Local*) Hint Resolve <- Z.log2_lt_pow2_alt : zarith.
- Local Hint Resolve eq_refl : zarith.
- Local Ltac wWToZ_t :=
- intros;
- try match goal with
- | [ |- ?wordToZ ?op = _ ]
- => let op' := head op in
- cbv [wordToZ op'] in *
- end;
- autorewrite with push_Zto_N push_Zof_N push_wordToN; try reflexivity.
- Local Ltac wWToZ_extra_t :=
- repeat first [ reflexivity
- | progress cbv [ModularBaseSystemListZOperations.neg ModularBaseSystemListZOperations.cmovne ModularBaseSystemListZOperations.cmovl] in *
- | progress break_match
- | progress fold_WordW_Z
- | progress intros
- | progress autorewrite with push_Zto_N push_Zof_N push_wordToN push_wordWToZ ].
-
- Local Notation bounds_statement wop Zop
- := ((0 <= Zop -> Z.log2 Zop < Z.of_nat bit_width -> wordWToZ wop = Zop)%Z).
- Local Notation bounds_statement_tuple wop Zop
- := ((HList.hlist (fun v => 0 <= v /\ Z.log2 v < Z.of_nat bit_width) Zop -> Tuple.map wordWToZ wop = Zop)%Z).
- Local Notation bounds_1statement wop Zop
- := (forall x,
- bounds_statement (wop x) (Zop (wordWToZ x))).
- Local Notation bounds_2statement wop Zop
- := (forall x y,
- bounds_statement (wop x y) (Zop (wordWToZ x) (wordWToZ y))).
- Local Notation bounds_4statement wop Zop
- := (forall x y z w,
- bounds_statement (wop x y z w) (Zop (wordWToZ x) (wordWToZ y) (wordWToZ z) (wordWToZ w))).
-
- Lemma wordWToZ_add : bounds_2statement add Z.add. Proof. wWToZ_t. Qed.
- Lemma wordWToZ_sub : bounds_2statement sub Z.sub. Proof. wWToZ_t. Qed.
- Lemma wordWToZ_mul : bounds_2statement mul Z.mul. Proof. wWToZ_t. Qed.
- Lemma wordWToZ_shl : bounds_2statement shl Z.shiftl.
- Proof.
- wWToZ_t; wWToZ_extra_t; unfold wordWToZ, wordBin.
- rewrite wordToN_NToWord_idempotent; [rewrite <- N2Z.inj_shiftl; reflexivity|].
- apply N2Z.inj_lt.
- rewrite N2Z.inj_shiftl.
- destruct (Z.lt_ge_cases 0 ((wordWToZ x) << (wordWToZ y)))%Z;
- [|eapply Z.le_lt_trans; [|apply N2Z.inj_lt, Npow2_gt0]; assumption].
- rewrite Npow2_N, N2Z.inj_pow, ?nat_N_Z.
- apply Z.log2_lt_pow2; assumption.
- Qed.
-
- Lemma wordWToZ_shr : bounds_2statement shr Z.shiftr.
- Proof.
- wWToZ_t; wWToZ_extra_t; unfold wordWToZ, wordBin.
- rewrite wordToN_NToWord_idempotent; [rewrite <- N2Z.inj_shiftr; reflexivity|].
- apply N2Z.inj_lt.
- rewrite N2Z.inj_shiftr.
- destruct (Z.lt_ge_cases 0 ((wordWToZ x) >> (wordWToZ y)))%Z;
- [|eapply Z.le_lt_trans; [|apply N2Z.inj_lt, Npow2_gt0]; assumption].
- rewrite Npow2_N, N2Z.inj_pow, nat_N_Z.
- apply Z.log2_lt_pow2; assumption.
- Qed.
-
- Lemma wordWToZ_land : bounds_2statement land Z.land.
- Proof. wWToZ_t. Qed.
- Lemma wordWToZ_lor : bounds_2statement lor Z.lor.
- Proof. wWToZ_t. Qed.
- Lemma wordWToZ_neg int_width : bounds_1statement (neg int_width) (ModularBaseSystemListZOperations.neg int_width).
- Proof. wWToZ_t; wWToZ_extra_t. Qed.
- Lemma wordWToZ_cmovne : bounds_4statement cmovne ModularBaseSystemListZOperations.cmovne.
- Proof. wWToZ_t; wWToZ_extra_t. Qed.
- Lemma wordWToZ_cmovle : bounds_4statement cmovle ModularBaseSystemListZOperations.cmovl.
- Proof. wWToZ_t; wWToZ_extra_t. Qed.
-
- Definition interp_base_type (t : base_type) : Type
- := match t with
- | _ => wordW
- end.
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := match f in op src dst return interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst with
- | OpConst _ v => fun _ => ZToWordW 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 => fst xy << snd xy
- | Shr _ => fun xy => fst xy >> snd xy
- | Land _ => fun xy => land (fst xy) (snd xy)
- | Lor _ => fun xy => lor (fst xy) (snd xy)
- | Neg _ int_width => fun x => neg int_width x
- | Cmovne _ => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovne x y z w
- | Cmovle _ => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovle x y z w
- | Cast _ _ => fun x => x
- end%wordW.
-
- Definition of_Z ty : Z.interp_base_type ty -> interp_base_type ty
- := match ty return Z.interp_base_type ty -> interp_base_type ty with
- | _ => ZToWordW
- end.
- Definition to_Z ty : interp_base_type ty -> Z.interp_base_type ty
- := match ty return interp_base_type ty -> Z.interp_base_type ty with
- | _ => wordWToZ
- end.
-
- Module Export Rewrites.
- Ltac wordW_util_arith := omega.
-
- Hint Rewrite wordWToZ_add using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_add using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_sub using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_sub using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_mul using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_mul using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_shl using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_shl using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_shr using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_shr using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_land using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_land using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_lor using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_lor using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_neg using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_neg using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_cmovne using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_cmovne using wordW_util_arith : pull_wordWToZ.
- Hint Rewrite wordWToZ_cmovle using wordW_util_arith : push_wordWToZ.
- Hint Rewrite <- wordWToZ_cmovle using wordW_util_arith : pull_wordWToZ.
- End Rewrites.
- End WordW.
-
- Module ZBounds.
- Export Bounds.
- Definition bounds := bounds.
- Bind Scope bounds_scope with bounds.
- Definition t := option bounds. (* TODO?: Separate out the bounds computation from the overflow computation? e.g., have [safety := in_bounds | overflow] and [t := bounds * safety]? *)
- Bind Scope bounds_scope with t.
- Local Coercion Z.of_nat : nat >-> Z.
- Definition wordWToBounds (x : WordW.wordW) : t
- := let v := WordW.wordWToZ x in Some {| lower := v ; upper := v |}.
- Definition SmartBuildBounds (l u : Z)
- := if ((0 <=? l) && (Z.log2 u <? WordW.bit_width))%Z%bool
- then Some {| lower := l ; upper := u |}
- else None.
- Definition t_map1 (f : bounds -> bounds) (x : t)
- := match x with
- | Some x
- => match f x with
- | Build_bounds l u
- => SmartBuildBounds l u
- end
- | _ => None
- end%Z.
- Definition t_map2 (f : bounds -> bounds -> bounds) (x y : t)
- := match x, y with
- | Some x, Some y
- => match f x y with
- | Build_bounds l u
- => SmartBuildBounds l u
- end
- | _, _ => None
- end%Z.
- Definition t_map4 (f : bounds -> bounds -> bounds -> bounds -> bounds) (x y z w : t)
- := match x, y, z, w with
- | Some x, Some y, Some z, Some w
- => match f x y z w with
- | Build_bounds l u
- => SmartBuildBounds l u
- end
- | _, _, _, _ => None
- end%Z.
- Definition add' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx + ly ; upper := ux + uy |}.
- Definition add : t -> t -> t := t_map2 add'.
- Definition sub' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx - uy ; upper := ux - ly |}.
- Definition sub : t -> t -> t := t_map2 sub'.
- Definition mul' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx * ly ; upper := ux * uy |}.
- Definition mul : t -> t -> t := t_map2 mul'.
- Definition shl' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx << ly ; upper := ux << uy |}.
- Definition shl : t -> t -> t := t_map2 shl'.
- Definition shr' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := lx >> uy ; upper := ux >> ly |}.
- Definition shr : t -> t -> t := t_map2 shr'.
- Definition land' : bounds -> bounds -> bounds
- := fun x y => let (lx, ux) := x in let (ly, uy) := y in {| lower := 0 ; upper := Z.min ux uy |}.
- Definition land : t -> t -> t := t_map2 land'.
- Definition lor' : bounds -> bounds -> 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 lor : t -> t -> t := t_map2 lor'.
- Definition neg' (int_width : Z) : bounds -> bounds
- := 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 := 0 ; upper := Z.ones int_width |}
- else {| lower := 0 ; upper := 0 |}.
- Definition neg (int_width : Z) : t -> t
- := fun v
- => if ((0 <=? int_width) && (int_width <=? WordW.bit_width))%Z%bool
- then t_map1 (neg' int_width) v
- else None.
- Definition cmovne' (r1 r2 : bounds) : bounds
- := 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 := t_map4 (fun _ _ => cmovne') x y r1 r2.
- Definition cmovle' (r1 r2 : bounds) : bounds
- := 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 := t_map4 (fun _ _ => cmovle') x y r1 r2.
-
- Module Export Notations.
- Delimit Scope bounds_scope with bounds.
- Notation "b[ l ~> u ]" := {| lower := l ; upper := u |} : bounds_scope.
- Infix "+" := add : bounds_scope.
- Infix "-" := sub : bounds_scope.
- Infix "*" := mul : bounds_scope.
- Infix "<<" := shl : bounds_scope.
- Infix ">>" := shr : bounds_scope.
- Infix "&'" := land : bounds_scope.
- End Notations.
-
- Definition interp_base_type (ty : base_type) : Type
- := option bounds.
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := match f in op src dst return interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst with
- | OpConst _ v => fun _ => SmartBuildBounds v 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 => fst xy << snd xy
- | Shr _ => fun xy => fst xy >> snd xy
- | Land _ => fun xy => land (fst xy) (snd xy)
- | Lor _ => fun xy => lor (fst xy) (snd xy)
- | Neg _ int_width => fun x => neg int_width x
- | Cmovne _ => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovne x y z w
- | Cmovle _ => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovle x y z w
- | Cast _ _ => fun x => x
- end%bounds.
-
- Definition of_wordW ty : WordW.interp_base_type ty -> interp_base_type ty
- := match ty return WordW.interp_base_type ty -> interp_base_type ty with
- | _ => wordWToBounds
- end.
-
- Ltac inversion_bounds :=
- let lower := (eval cbv [lower] in (fun x => lower x)) in
- let upper := (eval cbv [upper] in (fun y => upper y)) in
- repeat match goal with
- | [ H : _ = _ :> bounds |- _ ]
- => pose proof (f_equal lower H); pose proof (f_equal upper H); clear H;
- cbv beta iota in *
- | [ H : _ = _ :> t |- _ ]
- => unfold t in H; inversion_option
- end.
- End ZBounds.
-
- Module BoundedWordW.
- Export BoundedWord.
- Local Notation is_bounded_by value lower upper
- := ((0 <= lower /\ lower <= WordW.wordWToZ value <= upper /\ Z.log2 upper < Z.of_nat WordW.bit_width)%Z)
- (only parsing).
- Definition BoundedWord := BoundedWordGen WordW.wordW (fun value lower upper => is_bounded_by value lower upper).
- Bind Scope bounded_word_scope with BoundedWord.
- Definition t := option BoundedWord.
- Bind Scope bounded_word_scope with t.
- Local Coercion Z.of_nat : nat >-> Z.
-
- Ltac inversion_BoundedWord :=
- repeat match goal with
- | _ => progress subst
- | [ H : _ = _ :> BoundedWord |- _ ]
- => pose proof (f_equal lower H);
- pose proof (f_equal upper H);
- pose proof (f_equal value H);
- clear H
- end.
-
- Definition interp_base_type (ty : base_type)
- := LiftOption.interp_base_type' BoundedWord ty.
-
- Definition wordWToBoundedWord (x : WordW.wordW) : t.
- Proof.
- refine (let v := WordW.wordWToZ x in
- match Sumbool.sumbool_of_bool (0 <=? v)%Z, Sumbool.sumbool_of_bool (Z.log2 v <? Z.of_nat WordW.bit_width)%Z with
- | left Hl, left Hu
- => Some {| lower := WordW.wordWToZ x ; value := x ; upper := WordW.wordWToZ x |}
- | _, _ => None
- end).
- subst v.
- abstract (Z.ltb_to_lt; repeat split; (assumption || reflexivity)).
- Defined.
-
- Definition boundedWordToWordW (x : t) : WordW.wordW
- := match x with
- | Some x' => value x'
- | None => WordW.ZToWordW 0
- end.
-
- Definition of_wordW ty : WordW.interp_base_type ty -> interp_base_type ty
- := match ty return WordW.interp_base_type ty -> interp_base_type ty with
- | _ => wordWToBoundedWord
- end.
- Definition to_wordW ty : interp_base_type ty -> WordW.interp_base_type ty
- := match ty return interp_base_type ty -> WordW.interp_base_type ty with
- | _ => boundedWordToWordW
- end.
-
- (** XXX FIXME(jgross) This is going to break horribly if we need to support any types other than [Z] *)
- Definition to_wordW' ty : BoundedWord -> WordW.interp_base_type ty
- := match ty return BoundedWord -> WordW.interp_base_type ty with
- | _ => fun x => boundedWordToWordW (Some x)
- end.
-
- Definition to_Z' ty : BoundedWord -> Z.interp_base_type ty
- := fun x => WordW.to_Z _ (to_wordW' _ x).
-
- Definition of_Z ty : Z.interp_base_type ty -> interp_base_type ty
- := fun x => of_wordW _ (WordW.of_Z _ x).
- Definition to_Z ty : interp_base_type ty -> Z.interp_base_type ty
- := fun x => WordW.to_Z _ (to_wordW _ x).
-
- Definition BoundedWordToBounds (x : BoundedWord) : ZBounds.bounds
- := {| Bounds.lower := lower x ; Bounds.upper := upper x |}.
-
- Definition to_bounds' : t -> ZBounds.t
- := option_map BoundedWordToBounds.
-
- Definition to_bounds ty : interp_base_type ty -> ZBounds.interp_base_type ty
- := match ty return interp_base_type ty -> ZBounds.interp_base_type ty with
- | _ => to_bounds'
- end.
-
- Definition SmartBuildBoundedWord (v : Z) : t
- := if ((0 <=? v) && (Z.log2 v <? WordW.bit_width))%Z%bool
- then of_Z TZ v
- else None.
-
- Definition t_map1
- (opW : WordW.wordW -> WordW.wordW)
- (opB : ZBounds.t -> ZBounds.t)
- (pf : forall x l u,
- opB (Some (BoundedWordToBounds x))
- = Some {| Bounds.lower := l ; Bounds.upper := u |}
- -> let val := opW (value x) in
- is_bounded_by val l u)
- : t -> t
- := fun x : t
- => match x with
- | Some x
- => match opB (Some (BoundedWordToBounds x))
- as bop return opB (Some (BoundedWordToBounds x)) = bop -> t
- with
- | Some (Bounds.Build_bounds l u)
- => fun Heq => Some {| lower := l ; value := opW (value x) ; upper := u;
- in_bounds := pf _ _ _ Heq |}
- | None => fun _ => None
- end eq_refl
- | _ => None
- end.
-
- Definition t_map2
- (opW : WordW.wordW -> WordW.wordW -> WordW.wordW)
- (opB : ZBounds.t -> ZBounds.t -> ZBounds.t)
- (pf : forall x y l u,
- opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- = Some {| Bounds.lower := l ; Bounds.upper := u |}
- -> let val := opW (value x) (value y) in
- is_bounded_by val l u)
- : t -> t -> t
- := fun x y : t
- => match x, y with
- | Some x, Some y
- => match opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- as bop return opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) = bop -> t
- with
- | Some (Bounds.Build_bounds l u)
- => fun Heq => Some {| lower := l ; value := opW (value x) (value y) ; upper := u;
- in_bounds := pf _ _ _ _ Heq |}
- | None => fun _ => None
- end eq_refl
- | _, _ => None
- end.
-
- Definition t_map4
- (opW : WordW.wordW -> WordW.wordW -> WordW.wordW -> WordW.wordW -> WordW.wordW)
- (opB : ZBounds.t -> ZBounds.t -> ZBounds.t -> ZBounds.t -> ZBounds.t)
- (pf : forall x y z w l u,
- opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
- = Some {| Bounds.lower := l ; Bounds.upper := u |}
- -> let val := opW (value x) (value y) (value z) (value w) in
- is_bounded_by val l u)
- : t -> t -> t -> t -> t
- := fun x y z w : t
- => match x, y, z, w with
- | Some x, Some y, Some z, Some w
- => match opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
- as bop return opB _ _ _ _ = bop -> t
- with
- | Some (Bounds.Build_bounds l u)
- => fun Heq => Some {| lower := l ; value := opW (value x) (value y) (value z) (value w) ; upper := u;
- in_bounds := pf _ _ _ _ _ _ Heq |}
- | None => fun _ => None
- end eq_refl
- | _, _, _, _ => None
- end.
-
- Local Opaque WordW.bit_width.
- Hint Resolve Z.ones_nonneg : zarith.
- Local Ltac t_prestart :=
- repeat first [ match goal with
- | [ |- forall x l u, ?opB (Some (BoundedWordToBounds x)) = Some _ -> let val := ?opW (value x) in _ ]
- => let opB' := head opB in let opW' := head opW in progress (try unfold opB'; try unfold opW')
- | [ |- forall x y l u, ?opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) = Some _ -> let val := ?opW (value x) (value y) in _ ]
- => progress (try unfold opB; try unfold opW)
- | [ |- forall x y z w l u, ?opB _ _ _ _ = Some _ -> let val := ?opW (value x) (value y) (value z) (value w) in _ ]
- => progress (try unfold opB; try unfold opW)
- | [ |- appcontext[ZBounds.t_map1 ?op] ] => let op' := head op in unfold op'
- | [ |- appcontext[ZBounds.t_map2 ?op] ] => let op' := head op in unfold op'
- | [ |- appcontext[?op (Bounds.Build_bounds _ _)] ] => let op' := head op in unfold op'
- | [ |- appcontext[?op (Bounds.Build_bounds _ _) (Bounds.Build_bounds _ _)] ] => unfold op
- end
- | progress cbv [BoundedWordToBounds ZBounds.SmartBuildBounds cmovne cmovl ModularBaseSystemListZOperations.neg] in *
- | progress break_match
- | progress break_match_hyps
- | progress intros
- | progress subst
- | progress ZBounds.inversion_bounds
- | progress inversion_option
- | progress WordW.fold_WordW_Z
- | progress autorewrite with bool_congr_setoid in *
- | progress destruct_head' and
- | progress Z.ltb_to_lt
- | assumption
- | progress destruct_head' BoundedWord; simpl in *
- | progress autorewrite with push_wordWToZ
- | progress repeat apply conj
- | solve [ WordW.arith ]
- | progress destruct_head' or ].
- Local Ltac t_start :=
- repeat first [ progress t_prestart
- | match goal with
- | [ |- appcontext[Z.min ?x ?y] ]
- => apply (Z.min_case_strong x y)
- | [ |- appcontext[Z.max ?x ?y] ]
- => apply (Z.max_case_strong x y)
- end ].
-
- Local Hint Resolve WordW.bit_width_pos : zarith.
- Local Hint Extern 1 (Z.log2 _ < _)%Z => eapply Z.le_lt_trans; [ eapply Z.log2_le_mono; eassumption | eassumption ] : zarith.
- (* Local *) Hint Resolve <- Z.log2_lt_pow2_alt : zarith.
-
-
- Definition add : t -> t -> t.
- Proof.
- refine (t_map2 WordW.add ZBounds.add _);
- abstract (t_start; eapply add_valid_update; eauto).
- Defined.
-
- Definition sub : t -> t -> t.
- Proof.
- refine (t_map2 WordW.sub ZBounds.sub _);
- abstract (t_start; eapply sub_valid_update; eauto).
- Defined.
-
- Definition mul : t -> t -> t.
- Proof.
- refine (t_map2 WordW.mul ZBounds.mul _);
- abstract (t_start; eapply mul_valid_update; eauto).
- Defined.
-
- Definition land : t -> t -> t.
- Proof.
- refine (t_map2 WordW.land ZBounds.land _);
- abstract (t_prestart; eapply land_valid_update; eauto).
- Defined.
-
- Definition lor : t -> t -> t.
- Proof.
- refine (t_map2 WordW.lor ZBounds.lor _);
- abstract (t_prestart; eapply lor_valid_update; eauto).
- Defined.
-
- Definition shl : t -> t -> t.
- Proof.
- refine (t_map2 WordW.shl ZBounds.shl _);
- abstract (t_start; eapply shl_valid_update; eauto).
- Defined.
-
- Definition shr : t -> t -> t.
- Proof.
- refine (t_map2 WordW.shr ZBounds.shr _);
- abstract (t_start; eapply shr_valid_update; eauto).
- Defined.
-
- Definition neg (int_width : Z) : t -> t.
- Proof. refine (t_map1 (WordW.neg int_width) (ZBounds.neg int_width) _); abstract t_start. Defined.
-
- Definition cmovne : t -> t -> t -> t -> t.
- Proof. refine (t_map4 WordW.cmovne ZBounds.cmovne _); abstract t_start. Defined.
-
- Definition cmovle : t -> t -> t -> t -> t.
- Proof. refine (t_map4 WordW.cmovle ZBounds.cmovle _); abstract t_start. Defined.
-
- Local Notation unop_correct op opW opB :=
- (forall x v, op (Some x) = Some v
- -> value v = opW (value x)
- /\ Some (BoundedWordToBounds v) = opB (Some (BoundedWordToBounds x)))
- (only parsing).
-
- Local Notation binop_correct op opW opB :=
- (forall x y v, op (Some x) (Some y) = Some v
- -> value v = opW (value x) (value y)
- /\ Some (BoundedWordToBounds v) = opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)))
- (only parsing).
-
- Local Notation op4_correct op opW opB :=
- (forall x y z w v, op (Some x) (Some y) (Some z) (Some w) = Some v
- -> value v = opW (value x) (value y) (value z) (value w)
- /\ Some (BoundedWordToBounds v) = opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w)))
- (only parsing).
-
- Lemma t_map1_correct opW opB pf
- : unop_correct (t_map1 opW opB pf) opW opB.
- Proof.
- intros ?? H.
- unfold t_map1 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map2_correct opW opB pf
- : binop_correct (t_map2 opW opB pf) opW opB.
- Proof.
- intros ??? H.
- unfold t_map2 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map4_correct opW opB pf
- : op4_correct (t_map4 opW opB pf) opW opB.
- Proof.
- intros ????? H.
- unfold t_map4 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Local Notation unop_correct_None op opB :=
- (forall x, op (Some x) = None -> opB (Some (BoundedWordToBounds x)) = None)
- (only parsing).
-
- Local Notation binop_correct_None op opB :=
- (forall x y, op (Some x) (Some y) = None -> opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y)) = None)
- (only parsing).
-
- Local Notation op4_correct_None op opB :=
- (forall x y z w, op (Some x) (Some y) (Some z) (Some w) = None
- -> opB (Some (BoundedWordToBounds x)) (Some (BoundedWordToBounds y))
- (Some (BoundedWordToBounds z)) (Some (BoundedWordToBounds w))
- = None)
- (only parsing).
-
- Lemma t_map1_correct_None opW opB pf
- : unop_correct_None (t_map1 opW opB pf) opB.
- Proof.
- intros ? H.
- unfold t_map1 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map2_correct_None opW opB pf
- : binop_correct_None (t_map2 opW opB pf) opB.
- Proof.
- intros ?? H.
- unfold t_map2 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Lemma t_map4_correct_None opW opB pf
- : op4_correct_None (t_map4 opW opB pf) opB.
- Proof.
- intros ???? H.
- unfold t_map4 in H; convoy_destruct_in H; destruct_head' ZBounds.bounds;
- unfold BoundedWordToBounds in *;
- inversion_option; subst; simpl.
- eauto.
- Qed.
-
- Module Export Notations.
- Delimit Scope bounded_word_scope with bounded_word.
- Notation "b[ l ~> u ]" := {| lower := l ; upper := u |} : bounded_word_scope.
- Infix "+" := add : bounded_word_scope.
- Infix "-" := sub : bounded_word_scope.
- Infix "*" := mul : bounded_word_scope.
- Infix "<<" := shl : bounded_word_scope.
- Infix ">>" := shr : bounded_word_scope.
- Infix "&'" := land : bounded_word_scope.
- End Notations.
-
- Definition interp_op {src dst} (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
- := match f in op src dst return interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst with
- | OpConst _ v => fun _ => SmartBuildBoundedWord 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 => fst xy << snd xy
- | Shr _ => fun xy => fst xy >> snd xy
- | Land _ => fun xy => land (fst xy) (snd xy)
- | Lor _ => fun xy => lor (fst xy) (snd xy)
- | Neg _ int_width => fun x => neg int_width x
- | Cmovne _ => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovne x y z w
- | Cmovle _ => fun xyzw => let '(x, y, z, w) := eta4 xyzw in cmovle x y z w
- | Cast _ _ => fun x => x
- end%bounded_word.
- End BoundedWordW.
-
- Module ZBoundsTuple.
- Definition interp_flat_type (t : flat_type base_type)
- := LiftOption.interp_flat_type ZBounds.bounds t.
-
- Definition of_ZBounds {ty} : Syntax.interp_flat_type ZBounds.interp_base_type ty -> interp_flat_type ty
- := @LiftOption.of' ZBounds.bounds ty.
- Definition to_ZBounds {ty} : interp_flat_type ty -> Syntax.interp_flat_type ZBounds.interp_base_type ty
- := @LiftOption.to' ZBounds.bounds ty.
- End ZBoundsTuple.
-
- Module BoundedWordWTuple.
- Definition interp_flat_type (t : flat_type base_type)
- := LiftOption.interp_flat_type BoundedWordW.BoundedWord t.
-
- Definition of_BoundedWordW {ty} : Syntax.interp_flat_type BoundedWordW.interp_base_type ty -> interp_flat_type ty
- := @LiftOption.of' BoundedWordW.BoundedWord ty.
- Definition to_BoundedWordW {ty} : interp_flat_type ty -> Syntax.interp_flat_type BoundedWordW.interp_base_type ty
- := @LiftOption.to' BoundedWordW.BoundedWord ty.
- End BoundedWordWTuple.
-End InterpretationsGen.
diff --git a/src/Specific/GF25519BoundedCommon.v b/src/Specific/GF25519BoundedCommon.v
deleted file mode 100644
index a3914951e..000000000
--- a/src/Specific/GF25519BoundedCommon.v
+++ /dev/null
@@ -1,823 +0,0 @@
-Require Import Coq.Classes.Morphisms.
-Require Import Crypto.BaseSystem.
-Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
-Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
-Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
-Require Import Crypto.ModularArithmetic.ModularBaseSystem.
-Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
-Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
-Require Import Crypto.Specific.GF25519.
-Require Import Bedrock.Word Crypto.Util.WordUtil.
-Require Import Coq.Lists.List Crypto.Util.ListUtil.
-Require Import Crypto.Tactics.VerdiTactics.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Tuple.
-Require Import Crypto.Util.Tactics.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.Notations.
-Require Import Crypto.Util.Decidable.
-Require Import Crypto.Util.HList.
-Require Import Crypto.Util.Tuple.
-Require Import Crypto.Algebra.
-Import ListNotations.
-Require Import Coq.ZArith.ZArith Coq.ZArith.Zpower Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
-Local Open Scope Z.
-
-(* BEGIN common curve-specific definitions *)
-Definition bit_width : nat := Eval compute in Z.to_nat (GF25519.int_width).
-Local Notation b_of exp := (0, 2^exp + 2^(exp-3))%Z (only parsing). (* max is [(0, 2^(exp+2) + 2^exp + 2^(exp-1) + 2^(exp-3) + 2^(exp-4) + 2^(exp-5) + 2^(exp-6) + 2^(exp-10) + 2^(exp-12) + 2^(exp-13) + 2^(exp-14) + 2^(exp-15) + 2^(exp-17) + 2^(exp-23) + 2^(exp-24))%Z] *)
-Definition bounds_exp : tuple Z length_fe25519
- := Eval compute in
- Tuple.from_list length_fe25519 limb_widths eq_refl.
-Definition bounds : tuple (Z * Z) length_fe25519
- := Eval compute in
- Tuple.map (fun e => b_of e) bounds_exp.
-Definition wire_digit_bounds_exp : tuple Z (length wire_widths)
- := Eval compute in Tuple.from_list _ wire_widths eq_refl.
-Definition wire_digit_bounds : tuple (Z * Z) (length wire_widths)
- := Eval compute in Tuple.map (fun e => (0,2^e-1)%Z) wire_digit_bounds_exp.
-(* END common curve-specific definitions *)
-
-(* BEGIN aliases for word extraction *)
-Definition word64 := Word.word bit_width.
-Coercion word64ToZ (x : word64) : Z := Z.of_N (wordToN x).
-Coercion ZToWord64 (x : Z) : word64 := NToWord _ (Z.to_N x).
-Definition NToWord64 : N -> word64 := NToWord _.
-Definition word64ize (x : word64) : word64
- := Eval cbv [wordToN N.succ_double N.double] in NToWord64 (wordToN x).
-Definition w64eqb (x y : word64) := weqb x y.
-
-Global Arguments NToWord64 : simpl never.
-Arguments word64 : simpl never.
-Arguments bit_width : simpl never.
-Global Opaque word64.
-Global Opaque bit_width.
-
-(* END aliases for word extraction *)
-
-(* BEGIN basic types *)
-Module Type WordIsBounded.
- Parameter is_boundedT : forall (lower upper : Z), word64 -> bool.
- Parameter Build_is_boundedT : forall {lower upper} {proj_word : word64},
- andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true -> is_boundedT lower upper proj_word = true.
- Parameter project_is_boundedT : forall {lower upper} {proj_word : word64},
- is_boundedT lower upper proj_word = true -> andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true.
-End WordIsBounded.
-
-Module Import WordIsBoundedDefault : WordIsBounded.
- Definition is_boundedT : forall (lower upper : Z), word64 -> bool
- := fun lower upper proj_word => andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z.
- Definition Build_is_boundedT {lower upper} {proj_word : word64}
- : andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true -> is_boundedT lower upper proj_word = true
- := fun x => x.
- Definition project_is_boundedT {lower upper} {proj_word : word64}
- : is_boundedT lower upper proj_word = true -> andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true
- := fun x => x.
-End WordIsBoundedDefault.
-
-Definition bounded_word (lower upper : Z)
- := { proj_word : word64 | is_boundedT lower upper proj_word = true }.
-Local Notation word_of exp := (bounded_word (fst (b_of exp)) (snd (b_of exp))).
-Local Notation unbounded_word sz := (bounded_word 0 (2^sz-1)%Z).
-
-Local Opaque word64.
-Definition fe25519W := Eval cbv (*-[word64]*) in (tuple word64 length_fe25519).
-Definition wire_digitsW := Eval cbv (*-[word64]*) in (tuple word64 (length wire_widths)).
-Definition fe25519 :=
- Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
- hlist (fun e => word_of e) bounds_exp.
-Definition wire_digits :=
- Eval cbv -[bounded_word Z.pow Z.sub Z.add] in
- hlist (fun e => unbounded_word e) wire_digit_bounds_exp.
-
-Definition is_bounded_gen {n} (x : tuple Z n) (bounds : tuple (Z * Z) n) : bool
- := let res := Tuple.map2
- (fun bounds v =>
- let '(lower, upper) := bounds in
- (lower <=? v) && (v <=? upper))%bool%Z
- bounds x in
- List.fold_right andb true (Tuple.to_list _ res).
-
-Definition is_bounded (x : Specific.GF25519.fe25519) : bool
- := is_bounded_gen (n:=length_fe25519) x bounds.
-
-Definition wire_digits_is_bounded (x : Specific.GF25519.wire_digits) : bool
- := is_bounded_gen (n:=length wire_widths) x wire_digit_bounds.
-
-(* END basic types *)
-
-Section generic_destructuring.
- Fixpoint app_on' A n : forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f
- := match n return forall T (f : tuple' A n) (P : forall x : tuple' A n, T x), T f with
- | O => fun T v P => P v
- | S n' => fun T v P => let '(v, x) := v in app_on' A n' _ v (fun v => P (v, x))
- end.
- Definition app_on {A n} : forall {T} (f : tuple A n) (P : forall x : tuple A n, T x), T f
- := match n return forall T (f : tuple A n) (P : forall x : tuple A n, T x), T f with
- | O => fun T v P => P v
- | S n' => @app_on' A n'
- end.
- Lemma app_on'_correct {A n T} f (P : forall x : tuple' A n, T x) : app_on' A n T f P = P f.
- Proof.
- induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ (fun t => P (t, _))) ].
- Qed.
- Lemma app_on_correct {A n T} f (P : forall x : tuple A n, T x) : app_on f P = P f.
- Proof. destruct n; [ reflexivity | apply app_on'_correct ]. Qed.
-
- Fixpoint app_on_h' A F n : forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f
- := match n return forall ts T (f : @hlist' A n F ts) (P : forall x : @hlist' A n F ts, T x), T f with
- | O => fun ts T v P => P v
- | S n' => fun ts T v P => let '(v, x) := v in app_on_h' A F n' _ _ v (fun v => P (v, x))
- end.
- Definition app_on_h {A F n} : forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f
- := match n return forall ts T (f : @hlist A n F ts) (P : forall x : @hlist A n F ts, T x), T f with
- | O => fun ts T v P => P v
- | S n' => @app_on_h' A F n'
- end.
- Lemma app_on_h'_correct {A F n ts T} f P : @app_on_h' A F n ts T f P = P f.
- Proof.
- induction n; simpl in *; destruct_head' prod; [ reflexivity | exact (IHn _ _ _ (fun h => P (h, f))) ].
- Qed.
- Lemma app_on_h_correct {A} F {n} ts {T} f P : @app_on_h A F n ts T f P = P f.
- Proof. destruct n; [ reflexivity | apply app_on_h'_correct ]. Qed.
-
- Definition app_wire_digitsW_dep {A T} (P : forall x : tuple A (length wire_widths), T x)
- : forall (f : tuple A (length wire_widths)), T f
- := Eval compute in fun f => @app_on A (length wire_widths) T f P.
- Definition app_wire_digitsW {A T} (f : tuple A (length wire_widths)) (P : tuple A (length wire_widths) -> T)
- := Eval compute in @app_wire_digitsW_dep A (fun _ => T) P f.
- Definition app_fe25519W_dep {A T} (P : forall x : tuple A length_fe25519, T x)
- : forall (f : tuple A length_fe25519), T f
- := Eval compute in fun f => @app_on A length_fe25519 T f P.
- Definition app_fe25519W {A T} (f : tuple A length_fe25519) (P : tuple A length_fe25519 -> T)
- := Eval compute in @app_fe25519W_dep A (fun _ => T) P f.
- Definition app_fe25519_dep {T} (P : forall x : fe25519, T x)
- : forall f : fe25519, T f
- := Eval compute in fun f => @app_on_h _ (fun e => word_of e) length_fe25519 bounds_exp T f P.
- Definition app_fe25519 {T} (f : fe25519) (P : hlist (fun e => word_of e) bounds_exp -> T)
- := Eval compute in @app_fe25519_dep (fun _ => T) P f.
- Definition app_wire_digits_dep {T} (P : forall x : wire_digits, T x)
- : forall f : wire_digits, T f
- := Eval compute in fun f => @app_on_h _ (fun e => unbounded_word e) (length wire_widths) wire_digit_bounds_exp T f P.
- Definition app_wire_digits {T} (f : wire_digits) (P : hlist (fun e => unbounded_word e) wire_digit_bounds_exp -> T)
- := Eval compute in @app_wire_digits_dep (fun _ => T) P f.
-
- Definition app_wire_digitsW_dep_correct {A T} f P : @app_wire_digitsW_dep A T P f = P f
- := app_on_correct f P.
- Definition app_wire_digitsW_correct {A T} f P : @app_wire_digitsW A T f P = P f
- := @app_wire_digitsW_dep_correct A (fun _ => T) f P.
- Definition app_fe25519W_dep_correct {A T} f P : @app_fe25519W_dep A T P f = P f
- := app_on_correct f P.
- Definition app_fe25519W_correct {A T} f P : @app_fe25519W A T f P = P f
- := @app_fe25519W_dep_correct A (fun _ => T) f P.
- Definition app_fe25519_dep_correct {T} f P : @app_fe25519_dep T P f = P f
- := app_on_h_correct (fun e => word_of e) bounds_exp f P.
- Definition app_fe25519_correct {T} f P : @app_fe25519 T f P = P f
- := @app_fe25519_dep_correct (fun _ => T) f P.
- Definition app_wire_digits_dep_correct {T} f P : @app_wire_digits_dep T P f = P f
- := app_on_h_correct (fun e => unbounded_word e) wire_digit_bounds_exp f P.
- Definition app_wire_digits_correct {T} f P : @app_wire_digits T f P = P f
- := @app_wire_digits_dep_correct (fun _ => T) f P.
-
- Definition appify2 {T} (op : fe25519W -> fe25519W -> T) (f g : fe25519W) :=
- app_fe25519W f (fun f0 => (app_fe25519W g (fun g0 => op f0 g0))).
-
- Lemma appify2_correct : forall {T} op f g, @appify2 T op f g = op f g.
- Proof.
- intros. cbv [appify2].
- etransitivity; apply app_fe25519W_correct.
- Qed.
-
- Definition appify9 {T} (op : fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> T) (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe25519W) :=
- app_fe25519W x0 (fun x0' =>
- app_fe25519W x1 (fun x1' =>
- app_fe25519W x2 (fun x2' =>
- app_fe25519W x3 (fun x3' =>
- app_fe25519W x4 (fun x4' =>
- app_fe25519W x5 (fun x5' =>
- app_fe25519W x6 (fun x6' =>
- app_fe25519W x7 (fun x7' =>
- app_fe25519W x8 (fun x8' =>
- op x0' x1' x2' x3' x4' x5' x6' x7' x8'))))))))).
-
- Lemma appify9_correct : forall {T} op x0 x1 x2 x3 x4 x5 x6 x7 x8,
- @appify9 T op x0 x1 x2 x3 x4 x5 x6 x7 x8 = op x0 x1 x2 x3 x4 x5 x6 x7 x8.
- Proof.
- intros. cbv [appify9].
- repeat (etransitivity; [ apply app_fe25519W_correct | ]); reflexivity.
- Qed.
-End generic_destructuring.
-
-Definition eta_fe25519W_sig (x : fe25519W) : { v : fe25519W | v = x }.
-Proof.
- eexists; symmetry.
- repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
-Defined.
-Definition eta_fe25519W (x : fe25519W) : fe25519W
- := Eval cbv [proj1_sig eta_fe25519W_sig] in proj1_sig (eta_fe25519W_sig x).
-Definition eta_wire_digitsW_sig (x : wire_digitsW) : { v : wire_digitsW | v = x }.
-Proof.
- eexists; symmetry.
- repeat (etransitivity; [ apply surjective_pairing | apply f_equal2 ]); reflexivity.
-Defined.
-Definition eta_wire_digitsW (x : wire_digitsW) : wire_digitsW
- := Eval cbv [proj1_sig eta_wire_digitsW_sig] in proj1_sig (eta_wire_digitsW_sig x).
-
-Local Transparent word64.
-Lemma word64ize_id x : word64ize x = x.
-Proof. apply NToWord_wordToN. Qed.
-Local Opaque word64.
-
-Lemma word64eqb_Zeqb x y : (word64ToZ x =? word64ToZ y)%Z = w64eqb x y.
-Proof. apply wordeqb_Zeqb. Qed.
-
-Local Arguments Z.pow_pos !_ !_ / .
-Lemma word64ToZ_ZToWord64 x : 0 <= x < 2^Z.of_nat bit_width -> word64ToZ (ZToWord64 x) = x.
-Proof.
- intros; unfold word64ToZ, ZToWord64.
- rewrite wordToN_NToWord_idempotent, Z2N.id
- by (omega || apply N2Z.inj_lt; rewrite <- ?(N_nat_Z (Npow2 _)), ?Npow2_nat, ?Zpow_pow2, ?N2Z.id, ?Z2N.id, ?Z2Nat.id by omega; omega).
- reflexivity.
-Qed.
-Lemma ZToWord64_word64ToZ x : ZToWord64 (word64ToZ x) = x.
-Proof.
- intros; unfold word64ToZ, ZToWord64.
- rewrite N2Z.id, NToWord_wordToN; reflexivity.
-Qed.
-
-(* BEGIN precomputation. *)
-
-Definition proj_word {lower upper} (v : bounded_word lower upper) := Eval cbv [proj1_sig] in proj1_sig v.
-Definition word_bounded {lower upper} (v : bounded_word lower upper)
- : andb (lower <=? proj_word v)%Z (proj_word v <=? upper)%Z = true
- := project_is_boundedT (proj2_sig v).
-Definition Build_bounded_word' {lower upper} proj_word word_bounded : bounded_word lower upper
- := exist _ proj_word (Build_is_boundedT word_bounded).
-Arguments proj_word {_ _} _.
-Arguments word_bounded {_ _} _.
-Arguments Build_bounded_word' {_ _} _ _.
-Definition Build_bounded_word {lower upper} (proj_word : word64) (word_bounded : andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z = true)
- : bounded_word lower upper
- := Build_bounded_word'
- proj_word
- (match andb (lower <=? proj_word)%Z (proj_word <=? upper)%Z as b return b = true -> b = true with
- | true => fun _ => eq_refl
- | false => fun x => x
- end word_bounded).
-Lemma word_to_unbounded_helper {x e : nat} : (x < pow2 e)%nat -> (Z.of_nat e <= Z.of_nat bit_width)%Z -> ((0 <=? word64ToZ (ZToWord64 (Z.of_nat x))) && (word64ToZ (ZToWord64 (Z.of_nat x)) <=? 2 ^ (Z.of_nat e) - 1))%bool = true.
-Proof.
- rewrite pow2_id; intro H; apply Nat2Z.inj_lt in H; revert H.
- rewrite Z.pow_Zpow; simpl Z.of_nat.
- intros H H'.
- assert (2^Z.of_nat e <= 2^Z.of_nat bit_width) by auto with zarith.
- rewrite ?word64ToZ_ZToWord64 by omega.
- match goal with
- | [ |- context[andb ?x ?y] ]
- => destruct x eqn:?, y eqn:?; try reflexivity; Z.ltb_to_lt
- end;
- intros; omega.
-Qed.
-Definition word_to_unbounded_word {sz} (x : word sz) : (Z.of_nat sz <=? Z.of_nat bit_width)%Z = true -> unbounded_word (Z.of_nat sz).
-Proof.
- refine (fun pf => Build_bounded_word (Z.of_N (wordToN x)) _).
- abstract (rewrite wordToN_nat, nat_N_Z; Z.ltb_to_lt; apply (word_to_unbounded_helper (wordToNat_bound x)); simpl; omega).
-Defined.
-Definition word32_to_unbounded_word (x : word 32) : unbounded_word 32.
-Proof. apply (word_to_unbounded_word x); reflexivity. Defined.
-Definition word31_to_unbounded_word (x : word 31) : unbounded_word 31.
-Proof. apply (word_to_unbounded_word x); reflexivity. Defined.
-
-Local Opaque word64.
-Declare Reduction app_tuple_map := cbv [app_wire_digitsW app_fe25519W app_fe25519 HList.mapt HList.mapt' Tuple.map Tuple.map' fst snd on_tuple List.map List.app length_fe25519 List.length wire_widths Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' fst snd].
-Definition fe25519WToZ (x : fe25519W) : Specific.GF25519.fe25519
- := Eval app_tuple_map in
- app_fe25519W x (Tuple.map (fun v : word64 => v : Z)).
-Definition fe25519ZToW (x : Specific.GF25519.fe25519) : fe25519W
- := Eval app_tuple_map in
- app_fe25519W x (Tuple.map (fun v : Z => v : word64)).
-Definition wire_digitsWToZ (x : wire_digitsW) : Specific.GF25519.wire_digits
- := Eval app_tuple_map in
- app_wire_digitsW x (Tuple.map (fun v : word64 => v : Z)).
-Definition wire_digitsZToW (x : Specific.GF25519.wire_digits) : wire_digitsW
- := Eval app_tuple_map in
- app_wire_digitsW x (Tuple.map (fun v : Z => v : word64)).
-Definition fe25519W_word64ize (x : fe25519W) : fe25519W
- := Eval app_tuple_map in
- app_fe25519W x (Tuple.map word64ize).
-Definition wire_digitsW_word64ize (x : wire_digitsW) : wire_digitsW
- := Eval app_tuple_map in
- app_wire_digitsW x (Tuple.map word64ize).
-
-(** TODO: Turn this into a lemma to speed up proofs *)
-Ltac unfold_is_bounded_in' H :=
- lazymatch type of H with
- | andb _ _ = true
- => apply andb_prop in H;
- let H1 := fresh in
- let H2 := fresh in
- destruct H as [H1 H2];
- unfold_is_bounded_in' H1;
- unfold_is_bounded_in' H2
- | _ => idtac
- end.
-Ltac preunfold_is_bounded_in H :=
- unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe25519WToZ, wire_digitsWToZ in H;
- cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map Tuple.map' List.map fold_right List.rev List.app length_fe25519 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple] in H.
-Ltac unfold_is_bounded_in H :=
- preunfold_is_bounded_in H;
- unfold_is_bounded_in' H.
-
-Ltac preunfold_is_bounded :=
- unfold is_bounded, wire_digits_is_bounded, is_bounded_gen, fe25519WToZ, wire_digitsWToZ;
- cbv [to_list length bounds wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 Tuple.map Tuple.map' List.map fold_right List.rev List.app length_fe25519 List.length wire_widths HList.hlistP HList.hlistP' Tuple.on_tuple].
-
-Ltac unfold_is_bounded :=
- preunfold_is_bounded;
- repeat match goal with
- | [ |- andb _ _ = true ] => apply andb_true_intro
- | [ |- and _ _ ] => split
- end.
-
-Local Transparent bit_width.
-Definition Pow2_64 := Eval compute in 2^Z.of_nat bit_width.
-Definition unfold_Pow2_64 : 2^Z.of_nat bit_width = Pow2_64 := eq_refl.
-Local Opaque bit_width.
-
-Local Ltac prove_lt_bit_width :=
- rewrite unfold_Pow2_64; cbv [Pow2_64]; omega.
-
-Lemma fe25519ZToW_WToZ (x : fe25519W) : fe25519ZToW (fe25519WToZ x) = x.
-Proof.
- hnf in x; destruct_head' prod; cbv [fe25519WToZ fe25519ZToW].
- simpl; rewrite !ZToWord64_word64ToZ; reflexivity.
-Qed.
-
-Lemma fe25519WToZ_ZToW x : is_bounded x = true -> fe25519WToZ (fe25519ZToW x) = x.
-Proof.
- hnf in x; destruct_head' prod; cbv [fe25519WToZ fe25519ZToW].
- intro H.
- unfold_is_bounded_in H; destruct_head' and.
- Z.ltb_to_lt.
- simpl; rewrite !word64ToZ_ZToWord64 by prove_lt_bit_width.
- reflexivity.
-Qed.
-
-Lemma fe25519W_word64ize_id x : fe25519W_word64ize x = x.
-Proof.
- hnf in x; destruct_head' prod.
- cbv [fe25519W_word64ize];
- simpl; repeat apply f_equal2; apply word64ize_id.
-Qed.
-Lemma wire_digitsW_word64ize_id x : wire_digitsW_word64ize x = x.
-Proof.
- hnf in x; destruct_head' prod.
- cbv [wire_digitsW_word64ize];
- simpl; repeat apply f_equal2; apply word64ize_id.
-Qed.
-
-Definition uncurry_unop_fe25519W {T} (op : fe25519W -> T)
- := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length_fe25519) op.
-Definition curry_unop_fe25519W {T} op : fe25519W -> T
- := Eval cbv (*-[word64]*) in fun f => app_fe25519W f (Tuple.curry (n:=length_fe25519) op).
-Definition uncurry_binop_fe25519W {T} (op : fe25519W -> fe25519W -> T)
- := Eval cbv (*-[word64]*) in uncurry_unop_fe25519W (fun f => uncurry_unop_fe25519W (op f)).
-Definition curry_binop_fe25519W {T} op : fe25519W -> fe25519W -> T
- := Eval cbv (*-[word64]*) in appify2 (fun f => curry_unop_fe25519W (curry_unop_fe25519W op f)).
-
-Definition uncurry_unop_wire_digitsW {T} (op : wire_digitsW -> T)
- := Eval cbv (*-[word64]*) in Tuple.uncurry (n:=length wire_widths) op.
-Definition curry_unop_wire_digitsW {T} op : wire_digitsW -> T
- := Eval cbv (*-[word64]*) in fun f => app_wire_digitsW f (Tuple.curry (n:=length wire_widths) op).
-
-Definition uncurry_9op_fe25519W {T} (op : fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> T)
- := Eval cbv (*-[word64]*) in
- uncurry_unop_fe25519W (fun x0 =>
- uncurry_unop_fe25519W (fun x1 =>
- uncurry_unop_fe25519W (fun x2 =>
- uncurry_unop_fe25519W (fun x3 =>
- uncurry_unop_fe25519W (fun x4 =>
- uncurry_unop_fe25519W (fun x5 =>
- uncurry_unop_fe25519W (fun x6 =>
- uncurry_unop_fe25519W (fun x7 =>
- uncurry_unop_fe25519W (fun x8 =>
- op x0 x1 x2 x3 x4 x5 x6 x7 x8))))))))).
-Definition curry_9op_fe25519W {T} op : fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> fe25519W -> T
- := Eval cbv (*-[word64]*) in
- appify9 (fun x0 x1 x2 x3 x4 x5 x6 x7 x8
- => curry_unop_fe25519W (curry_unop_fe25519W (curry_unop_fe25519W (curry_unop_fe25519W (curry_unop_fe25519W (curry_unop_fe25519W (curry_unop_fe25519W (curry_unop_fe25519W (curry_unop_fe25519W op x0) x1) x2) x3) x4) x5) x6) x7) x8).
-
-Definition proj1_fe25519W (x : fe25519) : fe25519W
- := Eval app_tuple_map in
- app_fe25519 x (HList.mapt (fun _ => (@proj_word _ _))).
-Coercion proj1_fe25519 (x : fe25519) : Specific.GF25519.fe25519
- := fe25519WToZ (proj1_fe25519W x).
-
-Lemma is_bounded_proj1_fe25519 (x : fe25519) : is_bounded (proj1_fe25519 x) = true.
-Proof.
- revert x; refine (app_fe25519_dep _); intro x.
- hnf in x; destruct_head' prod; destruct_head' bounded_word.
- cbv [is_bounded proj1_fe25519 proj1_fe25519W fe25519WToZ to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word length_fe25519 is_bounded_gen].
- apply fold_right_andb_true_iff_fold_right_and_True.
- cbv [fold_right List.map].
- cbv beta in *.
- repeat split; auto using project_is_boundedT.
-Qed.
-
-Definition proj1_wire_digitsW (x : wire_digits) : wire_digitsW
- := app_wire_digits x (HList.mapt (fun _ => proj_word)).
-Coercion proj1_wire_digits (x : wire_digits) : Specific.GF25519.wire_digits
- := wire_digitsWToZ (proj1_wire_digitsW x).
-
-Lemma is_bounded_proj1_wire_digits (x : wire_digits) : wire_digits_is_bounded (proj1_wire_digits x) = true.
-Proof.
- revert x; refine (app_wire_digits_dep _); intro x.
- hnf in x; destruct_head' prod; destruct_head' bounded_word.
- cbv [wire_digits_is_bounded proj1_wire_digits proj1_wire_digitsW wire_digitsWToZ to_list length wire_digit_bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word is_bounded_gen wire_widths HList.mapt HList.mapt' app_wire_digits fst snd].
- apply fold_right_andb_true_iff_fold_right_and_True.
- cbv [fold_right List.map].
- cbv beta in *.
- repeat split; auto using project_is_boundedT.
-Qed.
-
-Local Ltac make_exist_W' x app_W_dep :=
- let H := fresh in
- revert x; refine (@app_W_dep _ _ _); intros x H;
- let x' := fresh in
- set (x' := x);
- cbv [tuple tuple' length_fe25519 List.length wire_widths] in x;
- destruct_head' prod;
- let rec do_refine v H :=
- first [ let v' := (eval cbv [snd fst] in (snd v)) in
- refine (_, Build_bounded_word v' _);
- [ do_refine (fst v) (proj2 H) | subst x'; abstract exact (proj1 H) ]
- | let v' := (eval cbv [snd fst] in v) in
- refine (Build_bounded_word v' _); subst x'; abstract exact (proj1 H) ] in
- let H' := constr:(proj1 (@fold_right_andb_true_iff_fold_right_and_True _) H) in
- let T := type of H' in
- let T := (eval cbv [id
- List.fold_right List.map List.length List.app ListUtil.map2 List.rev
- Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.map2 Tuple.on_tuple2
- fe25519 bounds fe25519WToZ length_fe25519
- wire_digits wire_digit_bounds wire_digitsWToZ wire_widths] in T) in
- let H' := constr:(H' : T) in
- let v := (eval unfold x' in x') in
- do_refine v H'.
-Local Ltac make_exist'' x exist_W ZToW :=
- let H := fresh in
- intro H; apply (exist_W (ZToW x));
- abstract (
- hnf in x; destruct_head' prod;
- let H' := fresh in
- pose proof H as H';
- unfold_is_bounded_in H;
- destruct_head' and; simpl in *;
- Z.ltb_to_lt;
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width;
- assumption
- ).
-Local Ltac make_exist' x app_W_dep exist'' exist_W ZToW :=
- let H := fresh in
- revert x; refine (@app_W_dep _ _ _); intros x H;
- let x' := fresh in
- set (x' := x) in *;
- cbv [tuple tuple' length_fe25519 List.length wire_widths] in x;
- destruct_head' prod;
- let rec do_refine v :=
- first [ let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word v)) in
- refine (Build_bounded_word v' _); subst x'; abstract exact (word_bounded v)
- | let v' := (eval cbv [exist_W ZToW exist'' proj_word Build_bounded_word Build_bounded_word' snd fst] in (proj_word (snd v))) in
- refine (_, Build_bounded_word v' _);
- [ do_refine (fst v) | subst x'; abstract exact (word_bounded (snd v)) ] ] in
- let v := (eval unfold x' in (exist'' x' H)) in
- do_refine v.
-
-Definition exist_fe25519W' (x : fe25519W) : is_bounded (fe25519WToZ x) = true -> fe25519.
-Proof. make_exist_W' x (@app_fe25519W_dep). Defined.
-Definition exist_fe25519W (x : fe25519W) : is_bounded (fe25519WToZ x) = true -> fe25519
- := Eval cbv [app_fe25519W_dep exist_fe25519W' fe25519ZToW] in exist_fe25519W' x.
-Definition exist_fe25519'' (x : Specific.GF25519.fe25519) : is_bounded x = true -> fe25519.
-Proof. make_exist'' x exist_fe25519W fe25519ZToW. Defined.
-Definition exist_fe25519' (x : Specific.GF25519.fe25519) : is_bounded x = true -> fe25519.
-Proof. make_exist' x (@app_fe25519W_dep) exist_fe25519'' exist_fe25519W fe25519ZToW. Defined.
-Definition exist_fe25519 (x : Specific.GF25519.fe25519) : is_bounded x = true -> fe25519
- := Eval cbv [exist_fe25519' exist_fe25519W exist_fe25519' app_fe25519 app_fe25519W_dep] in
- exist_fe25519' x.
-
-Lemma proj1_fe25519_exist_fe25519W x pf : proj1_fe25519 (exist_fe25519W x pf) = fe25519WToZ x.
-Proof. now hnf in x; destruct_head' prod. Qed.
-Lemma proj1_fe25519W_exist_fe25519 x pf : proj1_fe25519W (exist_fe25519 x pf) = fe25519ZToW x.
-Proof. now hnf in x; destruct_head' prod. Qed.
-Lemma proj1_fe25519_exist_fe25519 x pf : proj1_fe25519 (exist_fe25519 x pf) = x.
-Proof.
- hnf in x; destruct_head' prod.
- cbv [proj1_fe25519 exist_fe25519 proj1_fe25519W fe25519WToZ proj_word Build_bounded_word Build_bounded_word'].
- simpl.
- unfold_is_bounded_in pf.
- destruct_head' and.
- Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
- reflexivity.
-Qed.
-
-Definition exist_wire_digitsW' (x : wire_digitsW)
- : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits.
-Proof. make_exist_W' x (@app_wire_digitsW_dep). Defined.
-Definition exist_wire_digitsW (x : wire_digitsW)
- : wire_digits_is_bounded (wire_digitsWToZ x) = true -> wire_digits
- := Eval cbv [app_wire_digitsW_dep exist_wire_digitsW' wire_digitsZToW] in exist_wire_digitsW' x.
-Definition exist_wire_digits'' (x : Specific.GF25519.wire_digits)
- : wire_digits_is_bounded x = true -> wire_digits.
-Proof. make_exist'' x exist_wire_digitsW wire_digitsZToW. Defined.
-Definition exist_wire_digits' (x : Specific.GF25519.wire_digits)
- : wire_digits_is_bounded x = true -> wire_digits.
-Proof. make_exist' x (@app_wire_digitsW_dep) exist_wire_digits'' exist_wire_digitsW wire_digitsZToW. Defined.
-Definition exist_wire_digits (x : Specific.GF25519.wire_digits)
- : wire_digits_is_bounded x = true -> wire_digits
- := Eval cbv [exist_wire_digits' exist_wire_digitsW exist_wire_digits' app_wire_digits app_wire_digitsW_dep] in
- exist_wire_digits' x.
-
-Lemma proj1_wire_digits_exist_wire_digitsW x pf : proj1_wire_digits (exist_wire_digitsW x pf) = wire_digitsWToZ x.
-Proof. now hnf in x; destruct_head' prod. Qed.
-Lemma proj1_wire_digitsW_exist_wire_digits x pf : proj1_wire_digitsW (exist_wire_digits x pf) = wire_digitsZToW x.
-Proof. now hnf in x; destruct_head' prod. Qed.
-Lemma proj1_wire_digits_exist_wire_digits x pf : proj1_wire_digits (exist_wire_digits x pf) = x.
-Proof.
- hnf in x; destruct_head' prod.
- cbv [proj1_wire_digits exist_wire_digits proj1_wire_digitsW wire_digitsWToZ proj_word Build_bounded_word Build_bounded_word' app_wire_digits HList.mapt HList.mapt' length wire_widths fst snd].
- simpl.
- unfold_is_bounded_in pf.
- destruct_head' and.
- Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
- reflexivity.
-Qed.
-
-Module opt.
- Definition word64ToZ := Eval vm_compute in word64ToZ.
- Definition word64ToN := Eval vm_compute in @wordToN bit_width.
- Definition NToWord64 := Eval vm_compute in NToWord64.
- Definition bit_width := Eval vm_compute in bit_width.
- Definition Zleb := Eval cbv [Z.leb] in Z.leb.
- Definition andb := Eval vm_compute in andb.
- Definition word64ize := Eval vm_compute in word64ize.
-End opt.
-
-Local Transparent bit_width.
-Local Ltac do_change lem :=
- match lem with
- | context L[andb (?x <=? ?y)%Z (?y <=? ?z)]
- => let x' := (eval vm_compute in x) in
- let z' := (eval vm_compute in z) in
- lazymatch y with
- | word64ToZ (word64ize ?v)
- => let y' := constr:(opt.word64ToZ (opt.word64ize v)) in
- let L' := context L[andb (opt.Zleb x' y') (opt.Zleb y' z')] in
- do_change L'
- end
- | _ => lem
- end.
-Definition fe25519_word64ize (x : fe25519) : fe25519.
-Proof.
- set (x' := x).
- hnf in x; destruct_head' prod.
- let lem := constr:(exist_fe25519W (fe25519W_word64ize (proj1_fe25519W x'))) in
- let lem := (eval cbv [proj1_fe25519W x' fe25519W_word64ize proj_word exist_fe25519W Build_bounded_word' Build_bounded_word] in lem) in
- let lem := do_change lem in
- refine (lem _);
- change (is_bounded (fe25519WToZ (fe25519W_word64ize (proj1_fe25519W x'))) = true);
- abstract (rewrite fe25519W_word64ize_id; apply is_bounded_proj1_fe25519).
-Defined.
-Definition wire_digits_word64ize (x : wire_digits) : wire_digits.
-Proof.
- set (x' := x).
- hnf in x; destruct_head' prod.
- let lem := constr:(exist_wire_digitsW (wire_digitsW_word64ize (proj1_wire_digitsW x'))) in
- let lem := (eval cbv [proj1_wire_digitsW x' wire_digitsW_word64ize proj_word exist_wire_digitsW Build_bounded_word Build_bounded_word'] in lem) in
- let lem := do_change lem in
- let lem := (eval cbv [word64ize opt.word64ize andb Z.leb Z.compare CompOpp Pos.compare] in lem) in
- refine (lem _);
- change (wire_digits_is_bounded (wire_digitsWToZ (wire_digitsW_word64ize (proj1_wire_digitsW x'))) = true);
- abstract (rewrite wire_digitsW_word64ize_id; apply is_bounded_proj1_wire_digits).
-Defined.
-
-Lemma is_bounded_to_nth_default x (H : is_bounded x = true)
- : forall n : nat,
- (n < length limb_widths)%nat
- -> (0 <= nth_default 0 (Tuple.to_list length_fe25519 x) n <=
- snd (b_of (nth_default (-1) limb_widths n)))%Z.
-Proof.
- hnf in x; destruct_head' prod.
- unfold_is_bounded_in H; destruct_head' and.
- Z.ltb_to_lt.
- unfold nth_default; simpl.
- intros.
- repeat match goal with
- | [ |- context[nth_error _ ?x] ]
- => is_var x; destruct x; simpl
- end;
- omega.
-Qed.
-
-(* END precomputation *)
-
-(* Precompute constants *)
-
-Definition one' := Eval vm_compute in exist_fe25519 Specific.GF25519.one_ eq_refl.
-Definition one := Eval cbv [one' fe25519_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_word64ize one'.
-
-Definition zero' := Eval vm_compute in exist_fe25519 Specific.GF25519.zero_ eq_refl.
-Definition zero := Eval cbv [zero' fe25519_word64ize word64ize andb opt.word64ToZ opt.word64ize opt.Zleb Z.compare CompOpp Pos.compare Pos.compare_cont] in fe25519_word64ize zero'.
-
-Lemma fold_chain_opt_gen {A B} (F : A -> B) is_bounded ls id' op' id op chain
- (Hid_bounded : is_bounded (F id') = true)
- (Hid : id = F id')
- (Hop_bounded : forall x y, is_bounded (F x) = true
- -> is_bounded (F y) = true
- -> is_bounded (op (F x) (F y)) = true)
- (Hop : forall x y, is_bounded (F x) = true
- -> is_bounded (F y) = true
- -> op (F x) (F y) = F (op' x y))
- (Hls_bounded : forall n, is_bounded (F (nth_default id' ls n)) = true)
- : F (fold_chain_opt id' op' chain ls)
- = fold_chain_opt id op chain (List.map F ls)
- /\ is_bounded (F (fold_chain_opt id' op' chain ls)) = true.
-Proof.
- rewrite !fold_chain_opt_correct.
- revert dependent ls; induction chain as [|x xs IHxs]; intros.
- { pose proof (Hls_bounded 0%nat).
- destruct ls; simpl; split; trivial; congruence. }
- { destruct x; simpl; unfold Let_In; simpl.
- rewrite (fun ls pf => proj1 (IHxs ls pf)) at 1; simpl.
- { do 2 f_equal.
- rewrite <- Hop, Hid by auto.
- rewrite !map_nth_default_always.
- split; try reflexivity.
- apply (IHxs (_::_)).
- intros [|?]; autorewrite with simpl_nth_default; auto.
- rewrite <- Hop; auto. }
- { intros [|?]; simpl;
- autorewrite with simpl_nth_default; auto.
- rewrite <- Hop; auto. } }
-Qed.
-
-Lemma encode_bounded x : is_bounded (encode x) = true.
-Proof.
- pose proof (bounded_encode x).
- generalize dependent (encode x).
- intro t; compute in t; intros.
- destruct_head' prod.
- unfold Pow2Base.bounded in H.
- cbv [nth_default Tuple.to_list Tuple.to_list' List.length limb_widths params25519] in H.
- repeat match type of H with
- | context[nth_error (cons _ _) _]
- => let H' := fresh in
- pose proof (H O) as H'; specialize (fun i => H (S i)); simpl @nth_error in H, H';
- cbv beta iota in H'
- end.
- clear H.
- simpl in *.
- cbv [Z.pow_pos Z.mul Pos.mul Pos.iter nth_default nth_error value] in *.
- unfold is_bounded.
- apply fold_right_andb_true_iff_fold_right_and_True.
- cbv [is_bounded proj1_fe25519 to_list length bounds from_list from_list' map2 on_tuple2 to_list' ListUtil.map2 List.map List.rev List.app proj_word fold_right length_fe25519].
- repeat split; rewrite !Bool.andb_true_iff, !Z.leb_le; omega.
-Qed.
-
-Definition encode (x : F modulus) : fe25519
- := exist_fe25519 (encode x) (encode_bounded x).
-
-Definition decode (x : fe25519) : F modulus
- := ModularBaseSystem.decode (proj1_fe25519 x).
-
-Lemma proj1_fe25519_encode x
- : proj1_fe25519 (encode x) = ModularBaseSystem.encode x.
-Proof.
- cbv [encode].
- generalize (encode_bounded x); generalize (ModularBaseSystem.encode x).
- intros y pf; intros; hnf in y; destruct_head_hnf' prod.
- cbv [proj1_fe25519 exist_fe25519 proj1_fe25519W Build_bounded_word Build_bounded_word' fe25519WToZ proj_word].
- simpl.
- unfold_is_bounded_in pf.
- destruct_head' and.
- Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
- reflexivity.
-Qed.
-
-Lemma decode_exist_fe25519 x pf
- : decode (exist_fe25519 x pf) = ModularBaseSystem.decode x.
-Proof.
- hnf in x; destruct_head' prod.
- cbv [decode proj1_fe25519 exist_fe25519 proj1_fe25519W Build_bounded_word Build_bounded_word' fe25519WToZ proj_word].
- simpl.
- unfold_is_bounded_in pf.
- destruct_head' and.
- Z.ltb_to_lt.
- rewrite ?word64ToZ_ZToWord64 by prove_lt_bit_width.
- reflexivity.
-Qed.
-
-Definition div (f g : fe25519) : fe25519
- := exist_fe25519 (div (proj1_fe25519 f) (proj1_fe25519 g)) (encode_bounded _).
-
-Definition eq (f g : fe25519) : Prop := eq (proj1_fe25519 f) (proj1_fe25519 g).
-
-
-Notation in_op_correct_and_bounded k irop op
- := (((Tuple.map (n:=k) fe25519WToZ irop = op)
- /\ HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=k) fe25519WToZ irop))%type)
- (only parsing).
-
-(*Fixpoint inm_op_correct_and_bounded' (count_in count_out : nat)
- : forall (irop : Tower.tower_nd fe25519W (Tuple.tuple fe25519W count_out) count_in)
- (op : Tower.tower_nd GF25519.fe25519 (Tuple.tuple GF25519.fe25519 count_out) count_in)
- (cont : Prop -> Prop),
- Prop
- := match count_in return
- forall (irop : Tower.tower_nd fe25519W (Tuple.tuple fe25519W count_out) count_in)
- (op : Tower.tower_nd GF25519.fe25519 (Tuple.tuple GF25519.fe25519 count_out) count_in)
- (cont : Prop -> Prop),
- Prop
- with
- | O => fun irop op cont => cont (in_op_correct_and_bounded count_out irop op)
- | S n => fun irop op cont
- => forall x : fe25519W,
- @inm_op_correct_and_bounded'
- n count_out (irop x) (op (fe25519WToZ x))
- (fun P => cont (is_bounded (fe25519WToZ x) = true -> P))
- end.
-Definition inm_op_correct_and_bounded count_in count_out irop op
- := Eval cbv [inm_op_correct_and_bounded' Tower.tower_nd Tuple.tuple Tuple.tuple' HList.hlistP HList.hlistP'] in
- inm_op_correct_and_bounded' count_in count_out irop op (fun P => P).
-Fixpoint inm_op_correct_and_bounded_prefix' (count_in count_out : nat)
- : forall (irop : Tower.tower_nd fe25519W (Tuple.tuple fe25519W count_out) count_in)
- (op : Tower.tower_nd GF25519.fe25519 (Tuple.tuple GF25519.fe25519 count_out) count_in),
- Prop
- := match count_in return
- forall (irop : Tower.tower_nd fe25519W (Tuple.tuple fe25519W count_out) count_in)
- (op : Tower.tower_nd GF25519.fe25519 (Tuple.tuple GF25519.fe25519 count_out) count_in),
- Prop
- with
- | O => fun irop op => in_op_correct_and_bounded count_out irop op
- | S n => fun irop op
- => forall x : fe25519W,
- is_bounded (fe25519WToZ x) = true
- -> @inm_op_correct_and_bounded_prefix'
- n count_out (irop x) (op (fe25519WToZ x))
- end.
-Definition inm_op_correct_and_bounded_prefix count_in count_out irop op
- := inm_op_correct_and_bounded_prefix' count_in count_out irop op.
-
-Lemma inm_op_correct_and_bounded_iff_prefix' count_in count_out irop op
- (cont : Prop -> Prop)
- (cont_forall : forall T (P : T -> Prop), cont (forall x : T, P x) <-> forall x : T, cont (P x))
- : inm_op_correct_and_bounded' count_in count_out irop op cont <-> cont (inm_op_correct_and_bounded_prefix' count_in count_out irop op).
-Proof.
- revert dependent cont; induction count_in as [|count_in IHcount_in]; intros.
- { reflexivity. }
- { simpl.
- rewrite cont_forall.
- split; intros H' x; specialize (H' x);
- specialize (IHcount_in (irop x) (op (fe25519WToZ x)) (fun P => cont (is_bounded (fe25519WToZ x) = true -> P)));
- cbv beta in *;
- [ erewrite <- IHcount_in; [ assumption | .. ]
- | erewrite -> IHcount_in; [ assumption | .. ] ];
- clear IHcount_in.
- { intros; repeat setoid_rewrite cont_forall; split; eauto. }
- { intros; repeat setoid_rewrite cont_forall; split; eauto. } }
-Qed.
-
-Lemma inm_op_correct_and_bounded_iff_prefix count_in count_out irop op
- : inm_op_correct_and_bounded count_in count_out irop op <-> inm_op_correct_and_bounded_prefix count_in count_out irop op.
-Proof.
- apply (inm_op_correct_and_bounded_iff_prefix' count_in count_out irop op (fun P => P)).
- reflexivity.
-Qed.
-
-Definition inm_op_correct_and_bounded1 count_in irop op
- := Eval cbv [inm_op_correct_and_bounded Tuple.map Tuple.to_list Tuple.to_list' Tuple.from_list Tuple.from_list' Tuple.on_tuple List.map] in
- inm_op_correct_and_bounded count_in 1 irop op.*)
-Notation inm_op_correct_and_bounded n m irop op
- := ((forall x,
- HList.hlistP (fun v => is_bounded v = true) (Tuple.map (n:=n%nat%type) fe25519WToZ x)
- -> in_op_correct_and_bounded m (irop x) (op (Tuple.map (n:=n) fe25519WToZ x))))
- (only parsing).
-Notation ibinop_correct_and_bounded irop op := (inm_op_correct_and_bounded 2 1 irop op) (only parsing).
-Notation iunop_correct_and_bounded irop op := (inm_op_correct_and_bounded 1 1 irop op) (only parsing).
-Notation iunop_FEToZ_correct irop op
- := (forall x,
- is_bounded (fe25519WToZ x) = true
- -> word64ToZ (irop x) = op (fe25519WToZ x)) (only parsing).
-Notation iunop_FEToWire_correct_and_bounded irop op
- := (forall x,
- is_bounded (fe25519WToZ x) = true
- -> wire_digitsWToZ (irop x) = op (fe25519WToZ x)
- /\ wire_digits_is_bounded (wire_digitsWToZ (irop x)) = true) (only parsing).
-Notation iunop_WireToFE_correct_and_bounded irop op
- := (forall x,
- wire_digits_is_bounded (wire_digitsWToZ x) = true
- -> fe25519WToZ (irop x) = op (wire_digitsWToZ x)
- /\ is_bounded (fe25519WToZ (irop x)) = true) (only parsing).
-Notation i9top_correct_and_bounded k irop op := (inm_op_correct_and_bounded 9 k irop op) (only parsing).
-
-Notation prefreeze := GF25519.prefreeze.
diff --git a/src/Specific/GF25519Reflective/Common.v b/src/Specific/GF25519Reflective/Common.v
deleted file mode 100644
index f702216e5..000000000
--- a/src/Specific/GF25519Reflective/Common.v
+++ /dev/null
@@ -1,669 +0,0 @@
-Require Export Coq.ZArith.ZArith.
-Require Export Coq.Strings.String.
-Require Export Crypto.Specific.GF25519.
-Require Export Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Reify.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Wf.
-Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Tuple.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Eta.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Crypto.Reflection.Z.Interpretations64.Relations.
-Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
-Require Import Crypto.Reflection.Z.Reify.
-Require Export Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.Z.Syntax.Equality.
-Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.WfReflective.
-Require Import Crypto.Util.Curry.
-Require Import Crypto.Util.Tower.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.ListUtil.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.Tactics.
-Require Import Crypto.Util.Prod.
-Require Import Crypto.Util.Notations.
-
-Notation Expr := (Expr base_type op).
-
-Local Ltac make_type_from' T :=
- let T := (eval compute in T) in
- let rT := reify_type T in
- exact rT.
-Local Ltac make_type_from uncurried_op :=
- let T := (type of uncurried_op) in
- make_type_from' T.
-
-Definition fe25519T : flat_type base_type.
-Proof.
- let T := (eval compute in GF25519.fe25519) in
- let T := reify_flat_type T in
- exact T.
-Defined.
-Definition Expr_n_OpT (count_out : nat) : flat_type base_type
- := Eval cbv [Syntax.tuple Syntax.tuple' fe25519T] in
- Syntax.tuple fe25519T count_out.
-Definition Expr_nm_OpT (count_in count_out : nat) : type base_type
- := Eval cbv [Syntax.tuple Syntax.tuple' fe25519T Expr_n_OpT] in
- Arrow (Syntax.tuple fe25519T count_in) (Expr_n_OpT count_out).
-Definition ExprBinOpT : type base_type := Eval compute in Expr_nm_OpT 2 1.
-Definition ExprUnOpT : type base_type := Eval compute in Expr_nm_OpT 1 1.
-Definition ExprUnOpFEToZT : type base_type.
-Proof. make_type_from ge_modulus. Defined.
-Definition ExprUnOpWireToFET : type base_type.
-Proof. make_type_from unpack. Defined.
-Definition ExprUnOpFEToWireT : type base_type.
-Proof. make_type_from pack. Defined.
-Definition Expr4OpT : type base_type := Eval compute in Expr_nm_OpT 4 1.
-Definition Expr9_4OpT : type base_type := Eval compute in Expr_nm_OpT 9 4.
-Definition ExprArgT : flat_type base_type
- := Eval compute in domain ExprUnOpT.
-Definition ExprArgWireT : flat_type base_type
- := Eval compute in domain ExprUnOpWireToFET.
-Definition ExprZ : Type := Expr (Arrow Unit (Tbase TZ)).
-Definition ExprUnOpFEToZ : Type := Expr ExprUnOpFEToZT.
-Definition ExprUnOpWireToFE : Type := Expr ExprUnOpWireToFET.
-Definition ExprUnOpFEToWire : Type := Expr ExprUnOpFEToWireT.
-Definition Expr_nm_Op count_in count_out : Type := Expr (Expr_nm_OpT count_in count_out).
-Definition ExprBinOp : Type := Expr ExprBinOpT.
-Definition ExprUnOp : Type := Expr ExprUnOpT.
-Definition Expr4Op : Type := Expr Expr4OpT.
-Definition Expr9_4Op : Type := Expr Expr9_4OpT.
-Definition ExprArg : Type := Expr (Arrow Unit ExprArgT).
-Definition ExprArgWire : Type := Expr (Arrow Unit ExprArgWireT).
-Definition expr_nm_Op count_in count_out var : Type
- := expr base_type op (var:=var) (Expr_nm_OpT count_in count_out).
-Definition exprBinOp var : Type := expr base_type op (var:=var) ExprBinOpT.
-Definition exprUnOp var : Type := expr base_type op (var:=var) ExprUnOpT.
-Definition expr4Op var : Type := expr base_type op (var:=var) Expr4OpT.
-Definition expr9_4Op var : Type := expr base_type op (var:=var) Expr9_4OpT.
-Definition exprZ var : Type := expr base_type op (var:=var) (Arrow Unit (Tbase TZ)).
-Definition exprUnOpFEToZ var : Type := expr base_type op (var:=var) ExprUnOpFEToZT.
-Definition exprUnOpWireToFE var : Type := expr base_type op (var:=var) ExprUnOpWireToFET.
-Definition exprUnOpFEToWire var : Type := expr base_type op (var:=var) ExprUnOpFEToWireT.
-Definition exprArg var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgT).
-Definition exprArgWire var : Type := expr base_type op (var:=var) (Arrow Unit ExprArgWireT).
-
-Definition make_bound (x : Z * Z) : ZBounds.t
- := Some {| Bounds.lower := fst x ; Bounds.upper := snd x |}.
-
-Fixpoint Expr_nm_Op_bounds count_in count_out {struct count_in} : interp_flat_type ZBounds.interp_base_type (domain (Expr_nm_OpT count_in count_out))
- := match count_in return interp_flat_type _ (domain (Expr_nm_OpT count_in count_out)) with
- | 0 => tt
- | S n
- => let b := (Tuple.map make_bound bounds) in
- let bs := Expr_nm_Op_bounds n count_out in
- match n return interp_flat_type _ (domain (Expr_nm_OpT n _)) -> interp_flat_type _ (domain (Expr_nm_OpT (S n) _)) with
- | 0 => fun _ => b
- | S n' => fun bs => (bs, b)
- end bs
- end.
-Definition ExprBinOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprBinOpT)
- := Eval compute in Expr_nm_Op_bounds 2 1.
-Definition ExprUnOp_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpT)
- := Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToZ_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToZT)
- := Eval compute in Expr_nm_Op_bounds 1 1.
-Definition ExprUnOpFEToWire_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpFEToWireT)
- := Eval compute in Expr_nm_Op_bounds 1 1.
-Definition Expr4Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr4OpT)
- := Eval compute in Expr_nm_Op_bounds 4 1.
-Definition Expr9Op_bounds : interp_flat_type ZBounds.interp_base_type (domain Expr9_4OpT)
- := Eval compute in Expr_nm_Op_bounds 9 4.
-Definition ExprUnOpWireToFE_bounds : interp_flat_type ZBounds.interp_base_type (domain ExprUnOpWireToFET)
- := Tuple.map make_bound wire_digit_bounds.
-
-Definition interp_bexpr : ExprBinOp -> Specific.GF25519BoundedCommon.fe25519W * Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.fe25519W
- := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
-Definition interp_uexpr : ExprUnOp -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.fe25519W
- := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
-Definition interp_uexpr_FEToZ : ExprUnOpFEToZ -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.word64
- := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
-Definition interp_uexpr_FEToWire : ExprUnOpFEToWire -> Specific.GF25519BoundedCommon.fe25519W -> Specific.GF25519BoundedCommon.wire_digitsW
- := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
-Definition interp_uexpr_WireToFE : ExprUnOpWireToFE -> Specific.GF25519BoundedCommon.wire_digitsW -> Specific.GF25519BoundedCommon.fe25519W
- := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
-Definition interp_9_4expr : Expr9_4Op
- -> Tuple.tuple Specific.GF25519BoundedCommon.fe25519W 9
- -> Tuple.tuple Specific.GF25519BoundedCommon.fe25519W 4
- := fun e => interp_flat_type_eta (Interp (@WordW.interp_op) e).
-
-Notation binop_correct_and_bounded rop op
- := (ibinop_correct_and_bounded (interp_bexpr rop) (curry2 op)) (only parsing).
-Notation unop_correct_and_bounded rop op
- := (iunop_correct_and_bounded (interp_uexpr rop) op) (only parsing).
-Notation unop_FEToZ_correct rop op
- := (iunop_FEToZ_correct (interp_uexpr_FEToZ rop) op) (only parsing).
-Notation unop_FEToWire_correct_and_bounded rop op
- := (iunop_FEToWire_correct_and_bounded (interp_uexpr_FEToWire rop) op) (only parsing).
-Notation unop_WireToFE_correct_and_bounded rop op
- := (iunop_WireToFE_correct_and_bounded (interp_uexpr_WireToFE rop) op) (only parsing).
-Notation op9_4_correct_and_bounded rop op
- := (i9top_correct_and_bounded 4 (interp_9_4expr rop) op) (only parsing).
-
-Ltac rexpr_cbv :=
- lazymatch goal with
- | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?uncurry ?oper x } ]
- => let operf := head oper in
- let uncurryf := head uncurry in
- try cbv delta [T]; try cbv delta [oper];
- try cbv beta iota delta [uncurryf]
- | [ |- { rexpr | forall x, Interp _ (t:=?T) rexpr x = ?oper x } ]
- => let operf := head oper in
- try cbv delta [T]; try cbv delta [oper]
- end;
- cbv beta iota delta [interp_flat_type interp_base_type zero_ GF25519.fe25519 GF25519.wire_digits].
-
-Ltac reify_sig :=
- rexpr_cbv; eexists; Reify_rhs; reflexivity.
-
-Local Notation rexpr_sig T uncurried_op :=
- { rexprZ
- | forall x, Interp interp_op (t:=T) rexprZ x = uncurried_op x }
- (only parsing).
-
-Notation rexpr_binop_sig op := (rexpr_sig ExprBinOpT (curry2 op)) (only parsing).
-Notation rexpr_unop_sig op := (rexpr_sig ExprUnOpT op) (only parsing).
-Notation rexpr_unop_FEToZ_sig op := (rexpr_sig ExprUnOpFEToZT op) (only parsing).
-Notation rexpr_unop_FEToWire_sig op := (rexpr_sig ExprUnOpFEToWireT op) (only parsing).
-Notation rexpr_unop_WireToFE_sig op := (rexpr_sig ExprUnOpWireToFET op) (only parsing).
-Notation rexpr_9_4op_sig op := (rexpr_sig Expr9_4OpT op) (only parsing).
-
-Notation correct_and_bounded_genT ropW'v ropZ_sigv
- := (let ropW' := ropW'v in
- let ropZ_sig := ropZ_sigv in
- ropW' = proj1_sig ropZ_sig
- /\ interp_type_rel_pointwise Relations.related_Z (Interp (@BoundedWordW.interp_op) ropW') (Interp (@Z.interp_op) ropW')
- /\ interp_type_rel_pointwise Relations.related_bounds (Interp (@BoundedWordW.interp_op) ropW') (Interp (@ZBounds.interp_op) ropW')
- /\ interp_type_rel_pointwise Relations.related_wordW (Interp (@BoundedWordW.interp_op) ropW') (Interp (@WordW.interp_op) ropW'))
- (only parsing).
-
-Ltac app_tuples x y :=
- let tx := type of x in
- lazymatch (eval hnf in tx) with
- | prod _ _ => let xs := app_tuples (snd x) y in
- constr:((fst x, xs))
- | _ => constr:((x, y))
- end.
-
-Local Arguments Tuple.map2 : simpl never.
-Local Arguments Tuple.map : simpl never.
-(*
-Fixpoint args_to_bounded_helperT {n}
- (v : Tuple.tuple' WordW.wordW n)
- (bounds : Tuple.tuple' (Z * Z) n)
- (pf : List.fold_right
- andb true
- (Tuple.to_list
- _
- (Tuple.map2
- (n:=S n)
- (fun bounds v =>
- let '(lower, upper) := bounds in ((lower <=? v)%Z && (v <=? upper)%Z)%bool)
- bounds
- (Tuple.map (n:=S n) WordW.wordWToZ v))) = true)
- (res : Type)
- {struct n}
- : Type.
-Proof.
- refine (match n return (forall (v : Tuple.tuple' _ n) (bounds : Tuple.tuple' _ n),
- List.fold_right
- _ _ (Tuple.to_list
- _
- (Tuple.map2 (n:=S n) _ bounds (Tuple.map (n:=S n) _ v))) = true
- -> Type)
- with
- | 0 => fun v' bounds' pf0 => forall pf1 : (0 <= fst bounds' /\ Z.log2 (snd bounds') < Z.of_nat WordW.bit_width)%Z, res
- | S n' => fun v' bounds' pf0 => let t := _ in
- forall pf1 : (0 <= fst (snd bounds') /\ Z.log2 (snd (snd bounds')) < Z.of_nat WordW.bit_width)%Z, @args_to_bounded_helperT n' (fst v') (fst bounds') t res
- end v bounds pf).
- { clear -pf0.
- abstract (
- destruct v', bounds'; simpl @fst;
- rewrite Tuple.map_S in pf0;
- simpl in pf0;
- rewrite Tuple.map2_S in pf0;
- simpl @List.fold_right in *;
- rewrite Bool.andb_true_iff in pf0; tauto
- ). }
-Defined.
-
-Fixpoint args_to_bounded_helper {n} res
- {struct n}
- : forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res.
-Proof.
- refine match n return (forall v bounds pf, (Tuple.tuple' BoundedWordW.BoundedWord n -> res) -> @args_to_bounded_helperT n v bounds pf res) with
- | 0 => fun v bounds pf f pf' => f {| BoundedWord.lower := fst bounds ; BoundedWord.value := v ; BoundedWord.upper := snd bounds |}
- | S n'
- => fun v bounds pf f pf'
- => @args_to_bounded_helper
- n' res (fst v) (fst bounds) _
- (fun ts => f (ts, {| BoundedWord.lower := fst (snd bounds) ; BoundedWord.value := snd v ; BoundedWord.upper := snd (snd bounds) |}))
- end.
- { clear -pf pf'.
- unfold Tuple.map2, Tuple.map in pf; simpl in *.
- abstract (
- destruct bounds;
- simpl in *;
- rewrite !Bool.andb_true_iff in pf;
- destruct_head' and;
- Z.ltb_to_lt; auto
- ). }
- { simpl in *.
- clear -pf pf'.
- abstract (
- destruct bounds as [? [? ?] ], v; simpl in *;
- rewrite Tuple.map_S in pf; simpl in pf; rewrite Tuple.map2_S in pf;
- simpl in pf;
- rewrite !Bool.andb_true_iff in pf;
- destruct_head' and;
- Z.ltb_to_lt; auto
- ). }
-Defined.
-*)
-
-Definition assoc_right''
- := Eval cbv [Tuple.assoc_right' Tuple.rsnoc' fst snd] in @Tuple.assoc_right'.
-(*
-Definition args_to_bounded {n} v bounds pf
- := Eval cbv [args_to_bounded_helper assoc_right''] in
- @args_to_bounded_helper n _ v bounds pf (@assoc_right'' _ _).
-*)
-Local Ltac get_len T :=
- match (eval hnf in T) with
- | prod ?A ?B
- => let a := get_len A in
- let b := get_len B in
- (eval compute in (a + b)%nat)
- | _ => constr:(1%nat)
- end.
-
-Ltac assoc_right_tuple x so_far :=
- let t := type of x in
- lazymatch (eval hnf in t) with
- | prod _ _ => let so_far := assoc_right_tuple (snd x) so_far in
- assoc_right_tuple (fst x) so_far
- | _ => lazymatch so_far with
- | @None => x
- | _ => constr:((x, so_far))
- end
- end.
-
-(*
-Local Ltac make_args x :=
- let x' := fresh "x'" in
- compute in x |- *;
- let t := match type of x with @expr _ _ _ (Tflat ?t) => t end in
- let t' := match goal with |- @expr _ _ _ (Tflat ?t) => t end in
- refine (LetIn (invert_Return x) _);
- let x'' := fresh "x''" in
- intro x'';
- let xv := assoc_right_tuple x'' (@None) in
- refine (SmartVarf (xv : interp_flat_type _ t')).
-
-Local Ltac args_to_bounded x H :=
- let x' := fresh in
- set (x' := x);
- compute in x;
- let len := (let T := type of x in get_len T) in
- destruct_head' prod;
- let c := constr:(args_to_bounded (n:=pred len) x' _ H) in
- let bounds := lazymatch c with args_to_bounded _ ?bounds _ => bounds end in
- let c := (eval cbv [domain ExprUnOpT interp_flat_type args_to_bounded bounds pred fst snd] in c) in
- apply c; compute; clear;
- try abstract (
- repeat split;
- solve [ reflexivity
- | refine (fun v => match v with eq_refl => I end) ]
- ).
- *)
-
-Section gen.
- Definition bounds_are_good_gen
- {n : nat} (bounds : Tuple.tuple (Z * Z) n)
- := let res :=
- Tuple.map (fun bs => let '(lower, upper) := bs in ((0 <=? lower)%Z && (Z.log2 upper <? Z.of_nat WordW.bit_width)%Z)%bool) bounds
- in
- List.fold_right andb true (Tuple.to_list n res).
- Definition unop_args_to_bounded'
- (bs : Z * Z)
- (Hbs : bounds_are_good_gen (n:=1) bs = true)
- (x : word64)
- (H : is_bounded_gen (Tuple.map (fun v : word64 => (v : Z)) (n:=1) x) bs = true)
- : BoundedWordW.BoundedWord.
- Proof.
- refine {| BoundedWord.lower := fst bs ; BoundedWord.value := x ; BoundedWord.upper := snd bs |}.
- unfold bounds_are_good_gen, is_bounded_gen, Tuple.map, Tuple.map2 in *; simpl in *.
- abstract (
- destruct bs; Bool.split_andb; Z.ltb_to_lt; simpl;
- repeat apply conj; assumption
- ).
- Defined.
- Fixpoint n_op_args_to_bounded'
- n
- : forall (bs : Tuple.tuple' (Z * Z) n)
- (Hbs : bounds_are_good_gen (n:=S n) bs = true)
- (x : Tuple.tuple' word64 n)
- (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (n:=S n) (fun v : word64 => v : Z)) x) bs = true),
- interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Tbase TZ) n).
- Proof.
- destruct n as [|n']; simpl in *.
- { exact unop_args_to_bounded'. }
- { refine (fun bs Hbs x H
- => (@n_op_args_to_bounded' n' (fst bs) _ (fst x) _,
- @unop_args_to_bounded' (snd bs) _ (snd x) _));
- clear n_op_args_to_bounded';
- simpl in *;
- [ clear x H | clear Hbs | clear x H | clear Hbs ];
- unfold bounds_are_good_gen, is_bounded_gen in *;
- abstract (
- repeat first [ progress simpl in *
- | assumption
- | reflexivity
- | progress Bool.split_andb
- | progress destruct_head prod
- | match goal with
- | [ H : _ |- _ ] => progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep in H
- end
- | progress rewrite ?Tuple.map_S, ?Tuple.map2_S, ?Tuple.strip_eta_tuple'_dep
- | progress break_match_hyps
- | rewrite Bool.andb_true_iff; apply conj
- | unfold Tuple.map, Tuple.map', Tuple.map2; simpl; rewrite Bool.andb_true_iff; apply conj ]
- ). }
- Defined.
-
- Definition n_op_args_to_bounded
- n
- : forall (bs : Tuple.tuple (Z * Z) n)
- (Hbs : bounds_are_good_gen bs = true)
- (x : Tuple.tuple word64 n)
- (H : is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true),
- interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Tbase TZ) n)
- := match n with
- | 0 => fun _ _ _ _ => tt
- | S n' => @n_op_args_to_bounded' n'
- end.
-
- Fixpoint nm_op_args_to_bounded' n m
- (bs : Tuple.tuple (Z * Z) m)
- (Hbs : bounds_are_good_gen bs = true)
- : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple' (Tbase TZ) n))
- (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
- x),
- interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple' (Syntax.tuple (Tbase TZ) m) n)
- := match n with
- | 0 => @n_op_args_to_bounded m bs Hbs
- | S n' => fun x H
- => (@nm_op_args_to_bounded' n' m bs Hbs (fst x) (proj1 H),
- @n_op_args_to_bounded m bs Hbs (snd x) (proj2 H))
- end.
- Definition nm_op_args_to_bounded n m
- (bs : Tuple.tuple (Z * Z) m)
- (Hbs : bounds_are_good_gen bs = true)
- : forall (x : interp_flat_type (fun _ => Tuple.tuple word64 m) (Syntax.tuple (Tbase TZ) n))
- (H : SmartVarfPropMap (fun _ x => is_bounded_gen (Tuple.eta_tuple (Tuple.map (fun v : word64 => v : Z)) x) bs = true)
- x),
- interp_flat_type (fun _ => BoundedWordW.BoundedWord) (Syntax.tuple (Syntax.tuple (Tbase TZ) m) n)
- := match n with
- | 0 => fun _ _ => tt
- | S n' => @nm_op_args_to_bounded' n' m bs Hbs
- end.
-End gen.
-
-Local Ltac get_inner_len T :=
- lazymatch T with
- | (?T * _)%type => get_inner_len T
- | ?T => get_len T
- end.
-Local Ltac get_outer_len T :=
- lazymatch T with
- | (?A * ?B)%type => let a := get_outer_len A in
- let b := get_outer_len B in
- (eval compute in (a + b)%nat)
- | ?T => constr:(1%nat)
- end.
-Local Ltac args_to_bounded x H :=
- let t := type of x in
- let m := get_inner_len t in
- let n := get_outer_len t in
- let H := constr:(fun Hbs => @nm_op_args_to_bounded n m _ Hbs x H) in
- let H := (eval cbv beta in (H eq_refl)) in
- exact H.
-
-Definition binop_args_to_bounded (x : fe25519W * fe25519W)
- (H : is_bounded (fe25519WToZ (fst x)) = true)
- (H' : is_bounded (fe25519WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprBinOpT).
-Proof. args_to_bounded x (conj H H'). Defined.
-Definition unop_args_to_bounded (x : fe25519W) (H : is_bounded (fe25519WToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprUnOpT).
-Proof. args_to_bounded x H. Defined.
-Definition unopWireToFE_args_to_bounded (x : wire_digitsW) (H : wire_digits_is_bounded (wire_digitsWToZ x) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain ExprUnOpWireToFET).
-Proof. args_to_bounded x H. Defined.
-Definition op9_args_to_bounded (x : fe25519W * fe25519W * fe25519W * fe25519W * fe25519W * fe25519W * fe25519W * fe25519W * fe25519W)
- (H0 : is_bounded (fe25519WToZ (fst (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
- (H1 : is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst (fst (fst x))))))))) = true)
- (H2 : is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst (fst x)))))))) = true)
- (H3 : is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst x))))))) = true)
- (H4 : is_bounded (fe25519WToZ (snd (fst (fst (fst (fst x)))))) = true)
- (H5 : is_bounded (fe25519WToZ (snd (fst (fst (fst x))))) = true)
- (H6 : is_bounded (fe25519WToZ (snd (fst (fst x)))) = true)
- (H7 : is_bounded (fe25519WToZ (snd (fst x))) = true)
- (H8 : is_bounded (fe25519WToZ (snd x)) = true)
- : interp_flat_type (fun _ => BoundedWordW.BoundedWord) (domain Expr9_4OpT).
-Proof. args_to_bounded x (conj (conj (conj (conj (conj (conj (conj (conj H0 H1) H2) H3) H4) H5) H6) H7) H8). Defined.
-Local Ltac make_bounds_prop' bounds bounds' :=
- first [ refine (andb _ _);
- [ destruct bounds' as [bounds' _], bounds as [bounds _]
- | destruct bounds' as [_ bounds'], bounds as [_ bounds] ];
- try make_bounds_prop' bounds bounds'
- | exact (match bounds' with
- | Some bounds' => let (l, u) := bounds in
- let (l', u') := bounds' in
- ((l' <=? l) && (u <=? u'))%Z%bool
- | None => false
- end) ].
-Local Ltac make_bounds_prop bounds orig_bounds :=
- let bounds' := fresh "bounds'" in
- pose orig_bounds as bounds';
- make_bounds_prop' bounds bounds'.
-Definition unop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpT)) : bool.
-Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
-Definition binop_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprBinOpT)) : bool.
-Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
-Definition unopFEToWire_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToWireT)) : bool.
-Proof. make_bounds_prop bounds ExprUnOpWireToFE_bounds. Defined.
-Definition unopWireToFE_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpWireToFET)) : bool.
-Proof. make_bounds_prop bounds ExprUnOp_bounds. Defined.
-(* TODO FIXME(jgross?, andreser?): Is every function returning a single Z a boolean function? *)
-Definition unopFEToZ_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain ExprUnOpFEToZT)) : bool.
-Proof.
- refine (let (l, u) := bounds in ((0 <=? l) && (u <=? 1))%Z%bool).
-Defined.
-Definition op9_4_bounds_good (bounds : interp_flat_type (fun _ => ZBounds.bounds) (codomain Expr9_4OpT)) : bool.
-Proof. make_bounds_prop bounds Expr4Op_bounds. Defined.
-(*Definition ApplyUnOp {var} (f : exprUnOp var) : exprArg var -> exprArg var
- := fun x
- => LetIn (invert_Return (unop_make_args x))
- (fun k => invert_Return (Apply length_fe25519 f k)).
-Definition ApplyBinOp {var} (f : exprBinOp var) : exprArg var -> exprArg var -> exprArg var
- := fun x y
- => LetIn (invert_Return (unop_make_args x))
- (fun x'
- => LetIn (invert_Return (unop_make_args y))
- (fun y'
- => invert_Return (Apply length_fe25519
- (Apply length_fe25519
- f x') y'))).
-Definition ApplyUnOpFEToWire {var} (f : exprUnOpFEToWire var) : exprArg var -> exprArgWire var
- := fun x
- => LetIn (invert_Return (unop_make_args x))
- (fun k => invert_Return (Apply length_fe25519 f k)).
-Definition ApplyUnOpWireToFE {var} (f : exprUnOpWireToFE var) : exprArgWire var -> exprArg var
- := fun x
- => LetIn (invert_Return (unop_wire_make_args x))
- (fun k => invert_Return (Apply (List.length wire_widths) f k)).
-Definition ApplyUnOpFEToZ {var} (f : exprUnOpFEToZ var) : exprArg var -> exprZ var
- := fun x
- => LetIn (invert_Return (unop_make_args x))
- (fun k => invert_Return (Apply length_fe25519 f k)).
-*)
-
-(* FIXME TODO(jgross): This is a horrible tactic. We should unify the
- various kinds of correct and boundedness, and abstract in Gallina
- rather than Ltac *)
-Ltac t_correct_and_bounded ropZ_sig Hbounds H0 H1 args :=
- let Heq := fresh "Heq" in
- let Hbounds0 := fresh "Hbounds0" in
- let Hbounds1 := fresh "Hbounds1" in
- let Hbounds2 := fresh "Hbounds2" in
- pose proof (proj2_sig ropZ_sig) as Heq;
- cbv [interp_bexpr interp_uexpr interp_uexpr_FEToWire interp_uexpr_FEToZ interp_uexpr_WireToFE interp_9_4expr
- interp_flat_type_eta interp_flat_type_eta_gen
- curry_binop_fe25519W curry_unop_fe25519W curry_unop_wire_digitsW curry_9op_fe25519W
- curry_binop_fe25519 curry_unop_fe25519 curry_unop_wire_digits curry_9op_fe25519
- uncurry_binop_fe25519W uncurry_unop_fe25519W uncurry_unop_wire_digitsW uncurry_9op_fe25519W
- uncurry_binop_fe25519 uncurry_unop_fe25519 uncurry_unop_wire_digits uncurry_9op_fe25519
- ExprBinOpT ExprUnOpFEToWireT ExprUnOpT ExprUnOpFEToZT ExprUnOpWireToFET Expr9_4OpT Expr4OpT] in *;
- cbv zeta in *;
- simpl @fe25519WToZ; simpl @wire_digitsWToZ;
- rewrite <- Heq; clear Heq;
- destruct Hbounds as [Heq Hbounds];
- change interp_op with (@Z.interp_op) in *;
- change interp_base_type with (@Z.interp_base_type) in *;
- change word64 with WordW.wordW in *;
- rewrite <- Heq; clear Heq;
- destruct Hbounds as [ Hbounds0 [Hbounds1 Hbounds2] ];
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj_from_option2 WordW.to_Z pf Hbounds2 Hbounds0) as Hbounds_left;
- pose proof (fun pf => Relations.uncurry_interp_type_rel_pointwise_proj1_from_option2 Relations.related_wordW_boundsi' pf Hbounds1 Hbounds2) as Hbounds_right;
- specialize_by repeat first [ progress intros
- | progress unfold RelationClasses.Reflexive
- | reflexivity
- | assumption
- | progress destruct_head' base_type
- | progress destruct_head' BoundedWordW.BoundedWord
- | progress destruct_head' and
- | progress repeat apply conj ];
- specialize (Hbounds_left args H0);
- specialize (Hbounds_right args H0);
- cbv beta in *;
- lazymatch type of Hbounds_right with
- | match ?e with _ => _ end
- => lazymatch type of H1 with
- | match ?e' with _ => _ end
- => change e' with e in H1; destruct e eqn:?; [ | exfalso; assumption ]
- end
- end;
- repeat match goal with x := _ |- _ => subst x end;
- cbv [id
- binop_args_to_bounded unop_args_to_bounded unopWireToFE_args_to_bounded op9_args_to_bounded
- Tuple.map Tuple.map' Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list
- Relations.proj_eq_rel SmartVarfMap interp_flat_type smart_interp_flat_map domain fst snd BoundedWordW.to_wordW' BoundedWordW.boundedWordToWordW BoundedWord.value Relations.related_wordW_boundsi' Relations.related'_wordW_bounds Bounds.upper Bounds.lower codomain WordW.to_Z nm_op_args_to_bounded nm_op_args_to_bounded' n_op_args_to_bounded n_op_args_to_bounded' unop_args_to_bounded' Relations.interp_flat_type_rel_pointwise Relations.interp_flat_type_rel_pointwise_gen_Prop] in Hbounds_left, Hbounds_right;
- simpl @fst in Hbounds_left, Hbounds_right; simpl @snd in Hbounds_left, Hbounds_right;
- simpl @interp_flat_type in *;
- (let v := (eval unfold WordW.interp_base_type in (WordW.interp_base_type TZ)) in
- change (WordW.interp_base_type TZ) with v in *);
- cbv beta iota zeta in *;
- lazymatch goal with
- | [ |- fe25519WToZ ?x = _ /\ _ ]
- => generalize dependent x; intros
- | [ |- wire_digitsWToZ ?x = _ /\ _ ]
- => generalize dependent x; intros
- | [ |- (Tuple.map fe25519WToZ ?x = _) /\ _ ]
- => generalize dependent x; intros
- | [ |- ((Tuple.map fe25519WToZ ?x = _) * _)%type ]
- => generalize dependent x; intros
- | [ |- _ = _ ]
- => exact Hbounds_left
- end;
- cbv [interp_type interp_type_gen interp_type_gen_hetero interp_flat_type WordW.interp_base_type codomain] in *;
- destruct_head' prod;
- change word64ToZ with WordW.wordWToZ in *;
- (split; [ exact Hbounds_left | ]);
- cbv [interp_flat_type] in *;
- cbv [fst snd
- Tuple.map Tuple.map' Tuple.on_tuple Tuple.to_list Tuple.to_list' List.map Tuple.from_list Tuple.from_list Tuple.from_list'
- make_bound
- Datatypes.length wire_widths wire_digit_bounds PseudoMersenneBaseParams.limb_widths bounds
- binop_bounds_good unop_bounds_good unopFEToWire_bounds_good unopWireToFE_bounds_good unopFEToZ_bounds_good op9_4_bounds_good
- ExprUnOp_bounds ExprBinOp_bounds ExprUnOpFEToWire_bounds ExprUnOpFEToZ_bounds ExprUnOpWireToFE_bounds Expr9Op_bounds Expr4Op_bounds] in H1;
- destruct_head' ZBounds.bounds;
- unfold_is_bounded_in H1;
- simpl @fe25519WToZ; simpl @wire_digitsWToZ;
- destruct_head' and;
- Z.ltb_to_lt;
- change WordW.wordWToZ with word64ToZ in *;
- cbv [Tuple.map Tuple.map' HList.hlist Tuple.on_tuple Tuple.from_list Tuple.from_list' Tuple.to_list Tuple.to_list' List.map HList.hlist' fst snd fe25519WToZ HList.hlistP HList.hlistP'];
- cbv [WordW.bit_width BitSize64.bit_width Z.of_nat Pos.of_succ_nat Pos.succ] in *;
- repeat split; unfold_is_bounded;
- Z.ltb_to_lt;
- try omega; try reflexivity.
-
-Ltac rexpr_correct :=
- let ropW' := fresh in
- let ropZ_sig := fresh in
- intros ropW' ropZ_sig;
- let wf_ropW := fresh "wf_ropW" in
- assert (wf_ropW : Wf ropW') by (subst ropW' ropZ_sig; reflect_Wf base_type_eq_semidec_is_dec op_beq_bl);
- cbv zeta; repeat apply conj;
- [ vm_compute; reflexivity
- | apply @InterpWf;
- [ | apply wf_ropW ].. ];
- auto with interp_related.
-
-Notation rword_of_Z rexprZ_sig := (proj1_sig rexprZ_sig) (only parsing).
-
-Notation compute_bounds opW bounds
- := (Interp (@ZBounds.interp_op) opW bounds)
- (only parsing).
-
-Notation rexpr_wfT e := (Wf.Wf e) (only parsing).
-
-Ltac prove_rexpr_wfT
- := reflect_Wf Equality.base_type_eq_semidec_is_dec Equality.op_beq_bl.
-
-Module Export PrettyPrinting.
- (* We add [enlargen] to force [bounds_on] to be in [Type] in 8.4 and
- 8.5/8.6. Because [Set] is special and things break if
- [bounds_on] ends up in [Set] for reasons jgross hasn't bothered
- to debug. *)
- Inductive bounds_on := overflow | in_range (lower upper : Z) | enlargen (_ : Set).
-
- Inductive result := yes | no | borked.
-
- Definition ZBounds_to_bounds_on
- := fun (t : base_type) (x : ZBounds.interp_base_type t)
- => match x with
- | Some {| Bounds.lower := l ; Bounds.upper := u |}
- => in_range l u
- | None
- => overflow
- end.
-
- Fixpoint does_it_overflow {t} : interp_flat_type (fun t : base_type => bounds_on) t -> result
- := match t return interp_flat_type _ t -> result with
- | Tbase _ => fun v => match v with
- | overflow => yes
- | in_range _ _ => no
- | enlargen _ => borked
- end
- | Unit => fun _ => no
- | Prod x y => fun v => match @does_it_overflow _ (fst v), @does_it_overflow _ (snd v) with
- | no, no => no
- | yes, no | no, yes | yes, yes => yes
- | borked, _ | _, borked => borked
- end
- end.
-
- (** This gives a slightly easier to read version of the bounds *)
- Notation compute_bounds_for_display opW bounds
- := (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds)) (only parsing).
- Notation sanity_compute opW bounds
- := (does_it_overflow (SmartVarfMap ZBounds_to_bounds_on (compute_bounds opW bounds))) (only parsing).
- Notation sanity_check opW bounds
- := (eq_refl (sanity_compute opW bounds) <: no = no) (only parsing).
-End PrettyPrinting.
diff --git a/src/Specific/GF25519Reflective/Common9_4Op.v b/src/Specific/GF25519Reflective/Common9_4Op.v
deleted file mode 100644
index 71f75530e..000000000
--- a/src/Specific/GF25519Reflective/Common9_4Op.v
+++ /dev/null
@@ -1,102 +0,0 @@
-Require Export Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Util.Tactics.
-
-Local Opaque Interp.
-Local Notation eta_and x := (let (a, b) := x in a, let (a, b) := x in b) (only parsing).
-Lemma Expr9_4Op_correct_and_bounded
- ropW op (ropZ_sig : rexpr_9_4op_sig op)
- (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
- (H0 : forall x012345678
- (x012345678
- := (eta_fe25519W (fst (fst (fst (fst (fst (fst (fst (fst x012345678)))))))),
- eta_fe25519W (snd (fst (fst (fst (fst (fst (fst (fst x012345678)))))))),
- eta_fe25519W (snd (fst (fst (fst (fst (fst (fst x012345678))))))),
- eta_fe25519W (snd (fst (fst (fst (fst (fst x012345678)))))),
- eta_fe25519W (snd (fst (fst (fst (fst x012345678))))),
- eta_fe25519W (snd (fst (fst (fst x012345678)))),
- eta_fe25519W (snd (fst (fst x012345678))),
- eta_fe25519W (snd (fst x012345678)),
- eta_fe25519W (snd x012345678)))
- (Hx012345678
- : is_bounded (fe25519WToZ (fst (fst (fst (fst (fst (fst (fst (fst x012345678))))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst (fst (fst x012345678))))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst (fst x012345678)))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst x012345678))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst x012345678)))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst x012345678))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst x012345678)))) = true
- /\ is_bounded (fe25519WToZ (snd (fst x012345678))) = true
- /\ is_bounded (fe25519WToZ (snd x012345678)) = true),
- let (Hx0, Hx12345678) := (eta_and Hx012345678) in
- let (Hx1, Hx2345678) := (eta_and Hx12345678) in
- let (Hx2, Hx345678) := (eta_and Hx2345678) in
- let (Hx3, Hx45678) := (eta_and Hx345678) in
- let (Hx4, Hx5678) := (eta_and Hx45678) in
- let (Hx5, Hx678) := (eta_and Hx5678) in
- let (Hx6, Hx78) := (eta_and Hx678) in
- let (Hx7, Hx8) := (eta_and Hx78) in
- let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
- match LiftOption.of'
- (Interp (@BoundedWordW.interp_op) ropW
- (LiftOption.to' (Some args)))
- with
- | Some _ => True
- | None => False
- end)
- (H1 : forall x012345678
- (x012345678
- := (eta_fe25519W (fst (fst (fst (fst (fst (fst (fst (fst x012345678)))))))),
- eta_fe25519W (snd (fst (fst (fst (fst (fst (fst (fst x012345678)))))))),
- eta_fe25519W (snd (fst (fst (fst (fst (fst (fst x012345678))))))),
- eta_fe25519W (snd (fst (fst (fst (fst (fst x012345678)))))),
- eta_fe25519W (snd (fst (fst (fst (fst x012345678))))),
- eta_fe25519W (snd (fst (fst (fst x012345678)))),
- eta_fe25519W (snd (fst (fst x012345678))),
- eta_fe25519W (snd (fst x012345678)),
- eta_fe25519W (snd x012345678)))
- (Hx012345678
- : is_bounded (fe25519WToZ (fst (fst (fst (fst (fst (fst (fst (fst x012345678))))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst (fst (fst x012345678))))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst (fst x012345678)))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst (fst x012345678))))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst (fst x012345678)))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst (fst x012345678))))) = true
- /\ is_bounded (fe25519WToZ (snd (fst (fst x012345678)))) = true
- /\ is_bounded (fe25519WToZ (snd (fst x012345678))) = true
- /\ is_bounded (fe25519WToZ (snd x012345678)) = true),
- let (Hx0, Hx12345678) := (eta_and Hx012345678) in
- let (Hx1, Hx2345678) := (eta_and Hx12345678) in
- let (Hx2, Hx345678) := (eta_and Hx2345678) in
- let (Hx3, Hx45678) := (eta_and Hx345678) in
- let (Hx4, Hx5678) := (eta_and Hx45678) in
- let (Hx5, Hx678) := (eta_and Hx5678) in
- let (Hx6, Hx78) := (eta_and Hx678) in
- let (Hx7, Hx8) := (eta_and Hx78) in
- let args := op9_args_to_bounded x012345678 Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8 in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
- match LiftOption.of'
- (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
- with
- | Some bounds => op9_4_bounds_good bounds = true
- | None => False
- end)
- : op9_4_correct_and_bounded ropW op.
-Proof.
- intros xs Hxs.
- pose xs as xs'.
- compute in xs.
- destruct_head' prod.
- cbv [Tuple.map Tuple.on_tuple Tuple.to_list Tuple.to_list' fst snd List.map Tuple.from_list Tuple.from_list' HList.hlistP HList.hlistP'] in Hxs.
- pose Hxs as Hxs'.
- destruct Hxs as [ [ [ [ [ [ [ [ Hx0 Hx1 ] Hx2 ] Hx3 ] Hx4 ] Hx5 ] Hx6 ] Hx7 ] Hx8 ].
- specialize (H0 xs'
- (conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- specialize (H1 xs'
- (conj Hx0 (conj Hx1 (conj Hx2 (conj Hx3 (conj Hx4 (conj Hx5 (conj Hx6 (conj Hx7 Hx8))))))))).
- Time let args := constr:(op9_args_to_bounded xs' Hx0 Hx1 Hx2 Hx3 Hx4 Hx5 Hx6 Hx7 Hx8) in
- admit; t_correct_and_bounded ropZ_sig Hbounds H0 H1 args. (* On 8.6beta1, with ~2 GB RAM, Finished transaction in 46.56 secs (46.372u,0.14s) (successful) *)
-Admitted. (*Time Qed. (* On 8.6beta1, with ~4.3 GB RAM, Finished transaction in 67.652 secs (66.932u,0.64s) (successful) *)*)
diff --git a/src/Specific/GF25519Reflective/CommonBinOp.v b/src/Specific/GF25519Reflective/CommonBinOp.v
deleted file mode 100644
index bb69bfcb1..000000000
--- a/src/Specific/GF25519Reflective/CommonBinOp.v
+++ /dev/null
@@ -1,50 +0,0 @@
-Require Export Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Util.Tactics.
-
-Local Opaque Interp.
-Lemma ExprBinOp_correct_and_bounded
- ropW op (ropZ_sig : rexpr_binop_sig op)
- (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
- (H0 : forall xy
- (xy := (eta_fe25519W (fst xy), eta_fe25519W (snd xy)))
- (Hxy : is_bounded (fe25519WToZ (fst xy)) = true
- /\ is_bounded (fe25519WToZ (snd xy)) = true),
- let Hx := let (Hx, Hy) := Hxy in Hx in
- let Hy := let (Hx, Hy) := Hxy in Hy in
- let args := binop_args_to_bounded xy Hx Hy in
- match LiftOption.of'
- (Interp (@BoundedWordW.interp_op) ropW
- (LiftOption.to' (Some args)))
- with
- | Some _ => True
- | None => False
- end)
- (H1 : forall xy
- (xy := (eta_fe25519W (fst xy), eta_fe25519W (snd xy)))
- (Hxy : is_bounded (fe25519WToZ (fst xy)) = true
- /\ is_bounded (fe25519WToZ (snd xy)) = true),
- let Hx := let (Hx, Hy) := Hxy in Hx in
- let Hy := let (Hx, Hy) := Hxy in Hy in
- let args := binop_args_to_bounded (fst xy, snd xy) Hx Hy in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
- match LiftOption.of'
- (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
- with
- | Some bounds => binop_bounds_good bounds = true
- | None => False
- end)
- : binop_correct_and_bounded ropW op.
-Proof.
- intros xy HxHy.
- pose xy as xy'.
- compute in xy; destruct_head' prod.
- specialize (H0 xy' HxHy).
- specialize (H1 xy' HxHy).
- destruct HxHy as [Hx Hy].
- let args := constr:(binop_args_to_bounded xy' Hx Hy) in
- t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
-Qed.
diff --git a/src/Specific/GF25519Reflective/CommonUnOp.v b/src/Specific/GF25519Reflective/CommonUnOp.v
deleted file mode 100644
index 2c997051a..000000000
--- a/src/Specific/GF25519Reflective/CommonUnOp.v
+++ /dev/null
@@ -1,43 +0,0 @@
-Require Export Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Util.Tactics.
-
-Local Opaque Interp.
-Lemma ExprUnOp_correct_and_bounded
- ropW op (ropZ_sig : rexpr_unop_sig op)
- (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
- (H0 : forall x
- (x := eta_fe25519W x)
- (Hx : is_bounded (fe25519WToZ x) = true),
- let args := unop_args_to_bounded x Hx in
- match LiftOption.of'
- (Interp (@BoundedWordW.interp_op) ropW
- (LiftOption.to' (Some args)))
- with
- | Some _ => True
- | None => False
- end)
- (H1 : forall x
- (x := eta_fe25519W x)
- (Hx : is_bounded (fe25519WToZ x) = true),
- let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
- match LiftOption.of'
- (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
- with
- | Some bounds => unop_bounds_good bounds = true
- | None => False
- end)
- : unop_correct_and_bounded ropW op.
-Proof.
- intros x Hx.
- pose x as x'.
- hnf in x; destruct_head' prod.
- specialize (H0 x' Hx).
- specialize (H1 x' Hx).
- let args := constr:(unop_args_to_bounded x' Hx) in
- t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
-Qed.
diff --git a/src/Specific/GF25519Reflective/CommonUnOpFEToWire.v b/src/Specific/GF25519Reflective/CommonUnOpFEToWire.v
deleted file mode 100644
index 9230e440b..000000000
--- a/src/Specific/GF25519Reflective/CommonUnOpFEToWire.v
+++ /dev/null
@@ -1,43 +0,0 @@
-Require Export Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Util.Tactics.
-
-Local Opaque Interp.
-Lemma ExprUnOpFEToWire_correct_and_bounded
- ropW op (ropZ_sig : rexpr_unop_FEToWire_sig op)
- (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
- (H0 : forall x
- (x := eta_fe25519W x)
- (Hx : is_bounded (fe25519WToZ x) = true),
- let args := unop_args_to_bounded x Hx in
- match LiftOption.of'
- (Interp (@BoundedWordW.interp_op) ropW
- (LiftOption.to' (Some args)))
- with
- | Some _ => True
- | None => False
- end)
- (H1 : forall x
- (x := eta_fe25519W x)
- (Hx : is_bounded (fe25519WToZ x) = true),
- let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
- match LiftOption.of'
- (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
- with
- | Some bounds => unopFEToWire_bounds_good bounds = true
- | None => False
- end)
- : unop_FEToWire_correct_and_bounded ropW op.
-Proof.
- intros x Hx.
- pose x as x'.
- hnf in x; destruct_head' prod.
- specialize (H0 x' Hx).
- specialize (H1 x' Hx).
- let args := constr:(unop_args_to_bounded x' Hx) in
- t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
-Qed.
diff --git a/src/Specific/GF25519Reflective/CommonUnOpFEToZ.v b/src/Specific/GF25519Reflective/CommonUnOpFEToZ.v
deleted file mode 100644
index e8ae58b9a..000000000
--- a/src/Specific/GF25519Reflective/CommonUnOpFEToZ.v
+++ /dev/null
@@ -1,43 +0,0 @@
-Require Export Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Util.Tactics.
-
-Local Opaque Interp.
-Lemma ExprUnOpFEToZ_correct_and_bounded
- ropW op (ropZ_sig : rexpr_unop_FEToZ_sig op)
- (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
- (H0 : forall x
- (x := eta_fe25519W x)
- (Hx : is_bounded (fe25519WToZ x) = true),
- let args := unop_args_to_bounded x Hx in
- match LiftOption.of'
- (Interp (@BoundedWordW.interp_op) ropW
- (LiftOption.to' (Some args)))
- with
- | Some _ => True
- | None => False
- end)
- (H1 : forall x
- (x := eta_fe25519W x)
- (Hx : is_bounded (fe25519WToZ x) = true),
- let args := unop_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
- match LiftOption.of'
- (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
- with
- | Some bounds => unopFEToZ_bounds_good bounds = true
- | None => False
- end)
- : unop_FEToZ_correct ropW op.
-Proof.
- intros x Hx.
- pose x as x'.
- hnf in x; destruct_head' prod.
- specialize (H0 x' Hx).
- specialize (H1 x' Hx).
- let args := constr:(unop_args_to_bounded x' Hx) in
- t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
-Qed.
diff --git a/src/Specific/GF25519Reflective/CommonUnOpWireToFE.v b/src/Specific/GF25519Reflective/CommonUnOpWireToFE.v
deleted file mode 100644
index 3a861d03b..000000000
--- a/src/Specific/GF25519Reflective/CommonUnOpWireToFE.v
+++ /dev/null
@@ -1,43 +0,0 @@
-Require Export Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Util.Tactics.
-
-Local Opaque Interp.
-Lemma ExprUnOpWireToFE_correct_and_bounded
- ropW op (ropZ_sig : rexpr_unop_WireToFE_sig op)
- (Hbounds : correct_and_bounded_genT ropW ropZ_sig)
- (H0 : forall x
- (x := eta_wire_digitsW x)
- (Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
- let args := unopWireToFE_args_to_bounded x Hx in
- match LiftOption.of'
- (Interp (@BoundedWordW.interp_op) ropW
- (LiftOption.to' (Some args)))
- with
- | Some _ => True
- | None => False
- end)
- (H1 : forall x
- (x := eta_wire_digitsW x)
- (Hx : wire_digits_is_bounded (wire_digitsWToZ x) = true),
- let args := unopWireToFE_args_to_bounded x Hx in
- let x' := SmartVarfMap (fun _ : base_type => BoundedWordW.BoundedWordToBounds) args in
- match LiftOption.of'
- (Interp (@ZBounds.interp_op) ropW (LiftOption.to' (Some x')))
- with
- | Some bounds => unopWireToFE_bounds_good bounds = true
- | None => False
- end)
- : unop_WireToFE_correct_and_bounded ropW op.
-Proof.
- intros x Hx.
- pose x as x'.
- hnf in x; destruct_head' prod.
- specialize (H0 x' Hx).
- specialize (H1 x' Hx).
- let args := constr:(unopWireToFE_args_to_bounded x' Hx) in
- t_correct_and_bounded ropZ_sig Hbounds H0 H1 args.
-Qed.
diff --git a/src/Specific/GF25519Reflective/Reified.v b/src/Specific/GF25519Reflective/Reified.v
deleted file mode 100644
index aa50ad33f..000000000
--- a/src/Specific/GF25519Reflective/Reified.v
+++ /dev/null
@@ -1,13 +0,0 @@
-(** We split the reification up into separate files, one operation per
- file, so that it can run in parallel. *)
-Require Export Crypto.Specific.GF25519Reflective.Reified.Add.
-Require Export Crypto.Specific.GF25519Reflective.Reified.CarryAdd.
-Require Export Crypto.Specific.GF25519Reflective.Reified.Sub.
-Require Export Crypto.Specific.GF25519Reflective.Reified.CarrySub.
-Require Export Crypto.Specific.GF25519Reflective.Reified.Mul.
-Require Export Crypto.Specific.GF25519Reflective.Reified.Opp.
-Require Export Crypto.Specific.GF25519Reflective.Reified.CarryOpp.
-Require Export Crypto.Specific.GF25519Reflective.Reified.PreFreeze.
-Require Export Crypto.Specific.GF25519Reflective.Reified.GeModulus.
-Require Export Crypto.Specific.GF25519Reflective.Reified.Pack.
-Require Export Crypto.Specific.GF25519Reflective.Reified.Unpack.
diff --git a/src/Specific/GF25519Reflective/Reified/Add.v b/src/Specific/GF25519Reflective/Reified/Add.v
deleted file mode 100644
index f2700aca7..000000000
--- a/src/Specific/GF25519Reflective/Reified/Add.v
+++ /dev/null
@@ -1,12 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonBinOp.
-
-Definition raddZ_sig : rexpr_binop_sig add. Proof. reify_sig. Defined.
-Definition raddW := Eval vm_compute in rword_of_Z raddZ_sig.
-Lemma raddW_correct_and_bounded_gen : correct_and_bounded_genT raddW raddZ_sig.
-Proof. rexpr_correct. Qed.
-Definition radd_output_bounds := Eval vm_compute in compute_bounds raddW ExprBinOp_bounds.
-
-Local Open Scope string_scope.
-Compute ("Add", compute_bounds_for_display raddW ExprBinOp_bounds).
-Compute ("Add overflows? ", sanity_compute raddW ExprBinOp_bounds).
-Compute ("Add overflows (error if it does)? ", sanity_check raddW ExprBinOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/AddCoordinates.v b/src/Specific/GF25519Reflective/Reified/AddCoordinates.v
deleted file mode 100644
index 8ecce61f7..000000000
--- a/src/Specific/GF25519Reflective/Reified/AddCoordinates.v
+++ /dev/null
@@ -1,209 +0,0 @@
-Require Export Coq.ZArith.ZArith.
-Require Export Coq.Strings.String.
-Require Export Crypto.Specific.GF25519.
-Require Export Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Reify.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.Linearize.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Crypto.Reflection.Z.Interpretations64.Relations.
-Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
-Require Import Crypto.Reflection.Z.Reify.
-Require Export Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.LinearizeInterp.
-Require Import Crypto.Reflection.WfReflective.
-Require Import Crypto.Reflection.Eta.
-Require Import Crypto.Reflection.EtaInterp.
-Require Import Crypto.CompleteEdwardsCurve.ExtendedCoordinates.
-Require Import Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519Reflective.Reified.Add.
-Require Import Crypto.Specific.GF25519Reflective.Reified.Sub.
-Require Import Crypto.Specific.GF25519Reflective.Reified.Mul.
-Require Import Crypto.Specific.GF25519Reflective.Common9_4Op.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.HList.
-Require Import Crypto.Util.Tower.
-Require Import Crypto.Util.Tactics.
-Require Import Crypto.Util.Notations.
-Require Import Bedrock.Word.
-
-(* TODO: it would be nice if we didn't need to duplicate this *)
-Definition add_coordinates_gen (F:Type) (Fadd Fsub Fmul : F -> F -> F) (twice_d : F) (P1 P2 : F*F*F*F) (F4:Type) (parlet: F -> (F -> F4) -> F4) (pair4:F->F->F->F->F4) : F4 :=
- let (p, T1) := P1 in
- let (p0, Z1) := p in
- let (X1, Y1) := p0 in
- let (p1, T2) := P2 in
- let (p2, Z2) := p1 in
- let (X2, Y2) := p2 in
- parlet (Fmul (Fsub Y1 X1) (Fsub Y2 X2)) (fun A=>
- parlet (Fmul (Fadd Y1 X1) (Fadd Y2 X2)) (fun B=>
- parlet (Fmul (Fmul T1 twice_d) T2) (fun C=>
- parlet (Fmul Z1 (Fadd Z2 Z2)) (fun D=>
- parlet (Fsub B A) (fun E=>
- parlet (Fsub D C) (fun F0=>
- parlet (Fadd D C) (fun G=>
- parlet (Fadd B A) (fun H=>
- parlet (Fmul E F0) (fun X3=>
- parlet (Fmul G H) (fun Y3=>
- parlet (Fmul E H) (fun T3=>
- parlet (Fmul F0 G) (fun Z3=>
- pair4 X3 Y3 Z3 T3)))))))))))).
-
-Definition radd_coordinatesZ' var twice_d P1 P2
- := @add_coordinates_gen
- _
- (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
- (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
- (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
- twice_d
- P1 P2
- _
- (fun v f => LetIn v
- (fun k => f (SmartVarf k)))
- (fun x y z w => (x, y, z, w)%expr).
-
-Local Notation eta x := (fst x, snd x).
-
-Definition radd_coordinatesZ'' : Syntax.Expr _ _ _
- := Linearize
- (ExprEta
- (fun var
- => Abs (fun twice_d_P1_P2 : interp_flat_type _ (_ * ((_ * _ * _ * _) * (_ * _ * _ * _)))
- => let '(twice_d, ((P10, P11, P12, P13), (P20, P21, P22, P23)))
- := twice_d_P1_P2 in
- radd_coordinatesZ'
- var (SmartVarf twice_d)
- (SmartVarf P10, SmartVarf P11, SmartVarf P12, SmartVarf P13)
- (SmartVarf P20, SmartVarf P21, SmartVarf P22, SmartVarf P23)))).
-
-Definition add_coordinates
- := fun twice_d P10 P11 P12 P13 P20 P21 P22 P23
- => add_coordinates_gen
- _ add sub mul
- twice_d (P10, P11, P12, P13) (P20, P21, P22, P23) _ (fun x (f:_->(_*_*_*_)) => let y := x in f y) (fun x y z t => (x,y,z,t)).
-
-Definition uncurried_add_coordinates
- := fun twice_d_P1_P2
- => let twice_d := fst twice_d_P1_P2 in
- let (P1, P2) := eta (snd twice_d_P1_P2) in
- @add_coordinates_gen
- _ add sub mul
- twice_d P1 P2
- _
- (fun v f => dlet_nd k := v in f k)
- (fun x y z w => (x, y, z, w)).
-
-Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
- (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
- (only parsing).
-Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (forall x, rexpr_sigPf T uncurried_op rexprZ x)
- (only parsing).
-Local Notation rexpr_sig T uncurried_op :=
- { rexprZ | rexpr_sigP T uncurried_op rexprZ }
- (only parsing).
-
-Local Ltac fold_interpf' :=
- let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
- let k' := fresh in
- let H := fresh in
- pose k as k';
- assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- change k with k'; clearbody k'; subst k'.
-
-Local Ltac fold_interpf :=
- let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
- let k' := fresh in
- let H := fresh in
- pose k as k';
- assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- change k with k'; clearbody k'; subst k';
- fold_interpf'.
-
-Local Ltac repeat_step_interpf :=
- let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
- let k' := fresh in
- let H := fresh in
- pose k as k';
- assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold k'; change k with k'; unfold interpf_step);
- clearbody k'; subst k'.
-
-Lemma radd_coordinatesZ_sigP' : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- cbv [radd_coordinatesZ''].
- intro x; rewrite InterpLinearize, InterpExprEta.
- cbv [domain interp_flat_type] in x.
- destruct x as [twice_d [ [ [ [P10_ P11_] P12_] P13_] [ [ [P20_ P21_] P22_] P23_] ] ].
- repeat match goal with
- | [ H : prod _ _ |- _ ] => let H0 := fresh H in let H1 := fresh H in destruct H as [H0 H1]
- end.
- cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
- cbv [radd_coordinatesZ' uncurried_add_coordinates add_coordinates add_coordinates_gen SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
- repeat match goal with
- | [ |- appcontext[@proj1_sig ?A ?B ?v] ]
- => let k := fresh "f" in
- let k' := fresh "f" in
- let H := fresh in
- set (k := v);
- set (k' := @proj1_sig A B k);
- pose proof (proj2_sig k) as H;
- change (proj1_sig k) with k' in H;
- clearbody k'; clear k;
- cbv beta in *
- end.
- cbv [Interp Curry.curry2] in *.
- unfold interpf, interpf_step; fold_interpf.
- cbv beta iota delta [uncurried_add_coordinates add_coordinates_gen interp_flat_type interp_base_type GF25519.fe25519].
- Time
- abstract (
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros;
- [ cbv [Let_In] | ]
- end;
- repeat match goal with
- | _ => rewrite !interpf_invert_Abs
- | _ => rewrite_hyp !*
- | [ |- ?x = ?x ] => reflexivity
- | _ => rewrite <- !surjective_pairing
- end
- ).
-Time Defined.
-Lemma radd_coordinatesZ_sigP : rexpr_sigP _ uncurried_add_coordinates radd_coordinatesZ''.
-Proof.
- exact radd_coordinatesZ_sigP'.
-Qed.
-Definition radd_coordinatesZ_sig
- := exist (fun v => rexpr_sigP _ _ v) radd_coordinatesZ'' radd_coordinatesZ_sigP.
-
-Definition radd_coordinates_input_bounds
- := (ExprUnOp_bounds, ((ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds),
- (ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds, ExprUnOp_bounds))).
-
-Time Definition radd_coordinatesW := Eval vm_compute in rword_of_Z radd_coordinatesZ_sig.
-Lemma radd_coordinatesW_correct_and_bounded_gen : correct_and_bounded_genT radd_coordinatesW radd_coordinatesZ_sig.
-Proof. Time rexpr_correct. Time Qed.
-Definition radd_coordinates_output_bounds := Eval vm_compute in compute_bounds radd_coordinatesW radd_coordinates_input_bounds.
-
-Local Obligation Tactic := intros; vm_compute; constructor.
-
-(*
-Program Definition radd_coordinatesW_correct_and_bounded
- := Expr9_4Op_correct_and_bounded
- radd_coordinatesW uncurried_add_coordinates radd_coordinatesZ_sig radd_coordinatesW_correct_and_bounded_gen
- _ _.
- *)
-
-Local Open Scope string_scope.
-Compute ("Add_Coordinates", compute_bounds_for_display radd_coordinatesW radd_coordinates_input_bounds).
-Compute ("Add_Coordinates overflows? ", sanity_compute radd_coordinatesW radd_coordinates_input_bounds).
-Compute ("Add_Coordinates overflows (error if it does)? ", sanity_check radd_coordinatesW radd_coordinates_input_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/AddDisplay.log b/src/Specific/GF25519Reflective/Reified/AddDisplay.log
deleted file mode 100644
index 156f0286e..000000000
--- a/src/Specific/GF25519Reflective/Reified/AddDisplay.log
+++ /dev/null
@@ -1,45 +0,0 @@
-raddW =
-fun var : Syntax.base_type -> Type =>
-x x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17
- x18 : var Syntax.TZ,
-Tbase Syntax.TZ x19 = x + x9;
-Tbase Syntax.TZ x20 = x0 + x10;
-Tbase Syntax.TZ x21 = x1 + x11;
-Tbase Syntax.TZ x22 = x2 + x12;
-Tbase Syntax.TZ x23 = x3 + x13;
-Tbase Syntax.TZ x24 = x4 + x14;
-Tbase Syntax.TZ x25 = x5 + x15;
-Tbase Syntax.TZ x26 = x6 + x16;
-Tbase Syntax.TZ x27 = x7 + x17;
-Tbase Syntax.TZ x28 = x8 + x18;
-(Return x19, Return x20, Return x21, Return x22, Return x23,
-Return x24, Return x25, Return x26, Return x27, Return x28)
- : forall var : Syntax.base_type -> Type,
- expr Syntax.base_type Syntax.op
- (Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ)
-
-Argument scope is [function_scope]
diff --git a/src/Specific/GF25519Reflective/Reified/AddDisplay.v b/src/Specific/GF25519Reflective/Reified/AddDisplay.v
deleted file mode 100644
index 27b5519a3..000000000
--- a/src/Specific/GF25519Reflective/Reified/AddDisplay.v
+++ /dev/null
@@ -1,4 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.Reified.Add.
-Require Export Crypto.Reflection.Z.CNotations.
-
-Print raddW.
diff --git a/src/Specific/GF25519Reflective/Reified/AddJavaDisplay.log b/src/Specific/GF25519Reflective/Reified/AddJavaDisplay.log
deleted file mode 100644
index 156f0286e..000000000
--- a/src/Specific/GF25519Reflective/Reified/AddJavaDisplay.log
+++ /dev/null
@@ -1,45 +0,0 @@
-raddW =
-fun var : Syntax.base_type -> Type =>
-x x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17
- x18 : var Syntax.TZ,
-Tbase Syntax.TZ x19 = x + x9;
-Tbase Syntax.TZ x20 = x0 + x10;
-Tbase Syntax.TZ x21 = x1 + x11;
-Tbase Syntax.TZ x22 = x2 + x12;
-Tbase Syntax.TZ x23 = x3 + x13;
-Tbase Syntax.TZ x24 = x4 + x14;
-Tbase Syntax.TZ x25 = x5 + x15;
-Tbase Syntax.TZ x26 = x6 + x16;
-Tbase Syntax.TZ x27 = x7 + x17;
-Tbase Syntax.TZ x28 = x8 + x18;
-(Return x19, Return x20, Return x21, Return x22, Return x23,
-Return x24, Return x25, Return x26, Return x27, Return x28)
- : forall var : Syntax.base_type -> Type,
- expr Syntax.base_type Syntax.op
- (Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ)
-
-Argument scope is [function_scope]
diff --git a/src/Specific/GF25519Reflective/Reified/AddJavaDisplay.v b/src/Specific/GF25519Reflective/Reified/AddJavaDisplay.v
deleted file mode 100644
index 3daefb168..000000000
--- a/src/Specific/GF25519Reflective/Reified/AddJavaDisplay.v
+++ /dev/null
@@ -1,4 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.Reified.Add.
-Require Export Crypto.Reflection.Z.JavaNotations.
-
-Print raddW.
diff --git a/src/Specific/GF25519Reflective/Reified/CarryAdd.v b/src/Specific/GF25519Reflective/Reified/CarryAdd.v
deleted file mode 100644
index 6e05303e5..000000000
--- a/src/Specific/GF25519Reflective/Reified/CarryAdd.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonBinOp.
-
-Definition rcarry_addZ_sig : rexpr_binop_sig carry_add. Proof. reify_sig. Defined.
-Definition rcarry_addW := Eval vm_compute in rword_of_Z rcarry_addZ_sig.
-Lemma rcarry_addW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_addW rcarry_addZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rcarry_add_output_bounds := Eval vm_compute in compute_bounds rcarry_addW ExprBinOp_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition rcarry_addW_correct_and_bounded
- := ExprBinOp_correct_and_bounded
- rcarry_addW carry_add rcarry_addZ_sig rcarry_addW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("Carry_Add", compute_bounds_for_display rcarry_addW ExprBinOp_bounds).
-Compute ("Carry_Add overflows? ", sanity_compute rcarry_addW ExprBinOp_bounds).
-Compute ("Carry_Add overflows (error if it does)? ", sanity_check rcarry_addW ExprBinOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/CarryOpp.v b/src/Specific/GF25519Reflective/Reified/CarryOpp.v
deleted file mode 100644
index 52814dca9..000000000
--- a/src/Specific/GF25519Reflective/Reified/CarryOpp.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonUnOp.
-
-Definition rcarry_oppZ_sig : rexpr_unop_sig carry_opp. Proof. reify_sig. Defined.
-Definition rcarry_oppW := Eval vm_compute in rword_of_Z rcarry_oppZ_sig.
-Lemma rcarry_oppW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_oppW rcarry_oppZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rcarry_opp_output_bounds := Eval vm_compute in compute_bounds rcarry_oppW ExprUnOp_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition rcarry_oppW_correct_and_bounded
- := ExprUnOp_correct_and_bounded
- rcarry_oppW carry_opp rcarry_oppZ_sig rcarry_oppW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("Carry_Opp", compute_bounds_for_display rcarry_oppW ExprUnOp_bounds).
-Compute ("Carry_Opp overflows? ", sanity_compute rcarry_oppW ExprUnOp_bounds).
-Compute ("Carry_Opp overflows (error if it does)? ", sanity_check rcarry_oppW ExprUnOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/CarrySub.v b/src/Specific/GF25519Reflective/Reified/CarrySub.v
deleted file mode 100644
index 946a4ed67..000000000
--- a/src/Specific/GF25519Reflective/Reified/CarrySub.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonBinOp.
-
-Definition rcarry_subZ_sig : rexpr_binop_sig carry_sub. Proof. reify_sig. Defined.
-Definition rcarry_subW := Eval vm_compute in rword_of_Z rcarry_subZ_sig.
-Lemma rcarry_subW_correct_and_bounded_gen : correct_and_bounded_genT rcarry_subW rcarry_subZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rcarry_sub_output_bounds := Eval vm_compute in compute_bounds rcarry_subW ExprBinOp_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition rcarry_subW_correct_and_bounded
- := ExprBinOp_correct_and_bounded
- rcarry_subW carry_sub rcarry_subZ_sig rcarry_subW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("Carry_Sub", compute_bounds_for_display rcarry_subW ExprBinOp_bounds).
-Compute ("Carry_Sub overflows? ", sanity_compute rcarry_subW ExprBinOp_bounds).
-Compute ("Carry_Sub overflows (error if it does)? ", sanity_check rcarry_subW ExprBinOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/GeModulus.v b/src/Specific/GF25519Reflective/Reified/GeModulus.v
deleted file mode 100644
index 8bbcf0dc4..000000000
--- a/src/Specific/GF25519Reflective/Reified/GeModulus.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonUnOpFEToZ.
-
-Definition rge_modulusZ_sig : rexpr_unop_FEToZ_sig ge_modulus. Proof. reify_sig. Defined.
-Definition rge_modulusW := Eval vm_compute in rword_of_Z rge_modulusZ_sig.
-Lemma rge_modulusW_correct_and_bounded_gen : correct_and_bounded_genT rge_modulusW rge_modulusZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rge_modulus_output_bounds := Eval vm_compute in compute_bounds rge_modulusW ExprUnOpFEToZ_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition rge_modulusW_correct_and_bounded
- := ExprUnOpFEToZ_correct_and_bounded
- rge_modulusW ge_modulus rge_modulusZ_sig rge_modulusW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("Ge_Modulus", compute_bounds_for_display rge_modulusW ExprUnOpFEToZ_bounds).
-Compute ("Ge_Modulus overflows? ", sanity_compute rge_modulusW ExprUnOpFEToZ_bounds).
-Compute ("Ge_Modulus overflows (error if it does)? ", sanity_check rge_modulusW ExprUnOpFEToZ_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/LadderStep.v b/src/Specific/GF25519Reflective/Reified/LadderStep.v
deleted file mode 100644
index 1964fcaec..000000000
--- a/src/Specific/GF25519Reflective/Reified/LadderStep.v
+++ /dev/null
@@ -1,189 +0,0 @@
-Require Export Coq.ZArith.ZArith.
-Require Export Coq.Strings.String.
-Require Export Crypto.Specific.GF25519.
-Require Export Crypto.Specific.GF25519BoundedCommon.
-Require Import Crypto.Reflection.Reify.
-Require Import Crypto.Reflection.Syntax.
-Require Import Crypto.Reflection.SmartMap.
-Require Import Crypto.Reflection.Relations.
-Require Import Crypto.Reflection.ExprInversion.
-Require Import Crypto.Reflection.Linearize.
-Require Import Crypto.Reflection.Eta.
-Require Import Crypto.Reflection.EtaInterp.
-Require Import Crypto.Reflection.Z.Interpretations64.
-Require Crypto.Reflection.Z.Interpretations64.Relations.
-Require Import Crypto.Reflection.Z.Interpretations64.RelationsCombinations.
-Require Import Crypto.Reflection.Z.Reify.
-Require Export Crypto.Reflection.Z.Syntax.
-Require Import Crypto.Reflection.InterpWfRel.
-Require Import Crypto.Reflection.LinearizeInterp.
-Require Import Crypto.Reflection.WfReflective.
-Require Import Crypto.Spec.MxDH.
-Require Import Crypto.Specific.GF25519Reflective.Common.
-Require Import Crypto.Specific.GF25519Reflective.Reified.Add.
-Require Import Crypto.Specific.GF25519Reflective.Reified.Sub.
-Require Import Crypto.Specific.GF25519Reflective.Reified.Mul.
-Require Import Crypto.Specific.GF25519Reflective.Common9_4Op.
-Require Import Crypto.Util.LetIn.
-Require Import Crypto.Util.ZUtil.
-Require Import Crypto.Util.HList.
-Require Import Crypto.Util.Tower.
-Require Import Crypto.Util.Tactics.
-Require Import Crypto.Util.Notations.
-Require Import Bedrock.Word.
-
-Definition rladderstepZ' var (T:=_) (a24 x0 : T) P1 P2
- := @MxDH.ladderstep_gen
- _
- (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig raddZ_sig var)))
- (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rsubZ_sig var)))
- (fun x y => LetIn (Pair x y) (invert_Abs (proj1_sig rmulZ_sig var)))
- a24
- _
- (fun x y z w => (x, y, z, w)%expr)
- (fun v f => LetIn v
- (fun k => f (SmartVarf k)))
- x0
- P1 P2.
-
-Definition rladderstepZ'' : Syntax.Expr _ _ _
- := Linearize
- (ExprEta
- (fun var
- => Abs (fun a24_x0_P1_P2 : interp_flat_type _ (_ * _ * ((_ * _) * (_ * _)))
- => let '(a24, x0, ((P10, P11), (P20, P21)))
- := a24_x0_P1_P2 in
- rladderstepZ'
- var (SmartVarf a24) (SmartVarf x0)
- (SmartVarf P10, SmartVarf P11)
- (SmartVarf P20, SmartVarf P21)))).
-
-Local Notation eta x := (fst x, snd x).
-
-Definition ladderstep_other_assoc {F Fadd Fsub Fmul} a24 (X1:F) (P1 P2:F*F) : F*F*F*F :=
- Eval cbv beta delta [MxDH.ladderstep_gen] in
- @MxDH.ladderstep_gen
- F Fadd Fsub Fmul a24
- (F*F*F*F)
- (fun X3 Y3 Z3 T3 => (X3, Y3, Z3, T3))
- (fun x f => dlet y := x in f y)
- X1 P1 P2.
-
-Definition uncurried_ladderstep
- := fun (a24_x0_P1_P2 : _ * _ * ((_ * _) * (_ * _)))
- => let a24 := fst (fst a24_x0_P1_P2) in
- let x0 := snd (fst a24_x0_P1_P2) in
- let '(P1, P2) := eta (snd a24_x0_P1_P2) in
- let '((P10, P11), (P20, P21)) := (eta P1, eta P2) in
- @ladderstep_other_assoc
- _ add sub mul
- a24 x0 (P10, P11) (P20, P21).
-
-Local Notation rexpr_sigPf T uncurried_op rexprZ x :=
- (Interp interp_op (t:=T) rexprZ x = uncurried_op x)
- (only parsing).
-Local Notation rexpr_sigP T uncurried_op rexprZ :=
- (forall x, rexpr_sigPf T uncurried_op rexprZ x)
- (only parsing).
-Local Notation rexpr_sig T uncurried_op :=
- { rexprZ | rexpr_sigP T uncurried_op rexprZ }
- (only parsing).
-
-Local Ltac fold_interpf' :=
- let k := (eval unfold interpf, interpf_step in (@interpf base_type interp_base_type op interp_op)) in
- let k' := fresh in
- let H := fresh in
- pose k as k';
- assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- change k with k'; clearbody k'; subst k'.
-
-Local Ltac fold_interpf :=
- let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
- let k' := fresh in
- let H := fresh in
- pose k as k';
- assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- change k with k'; clearbody k'; subst k';
- fold_interpf'.
-
-Local Ltac repeat_step_interpf :=
- let k := (eval unfold interpf in (@interpf base_type interp_base_type op interp_op)) in
- let k' := fresh in
- let H := fresh in
- pose k as k';
- assert (H : @interpf base_type interp_base_type op interp_op = k') by reflexivity;
- repeat (unfold interpf_step at 1; change k with k' at 1);
- clearbody k'; subst k'.
-
-Lemma rladderstepZ_sigP' : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
-Proof.
- cbv [rladderstepZ''].
- intro x; rewrite InterpLinearize, InterpExprEta.
- cbv [domain interp_flat_type interp_base_type] in x.
- destruct_head' prod.
- cbv [invert_Abs domain codomain Interp interp SmartVarf smart_interp_flat_map fst snd].
- cbv [rladderstepZ' MxDH.ladderstep_gen uncurried_ladderstep SmartVarf smart_interp_flat_map]; simpl @fst; simpl @snd.
- repeat match goal with
- | [ |- appcontext[@proj1_sig ?A ?B ?v] ]
- => let k := fresh "f" in
- let k' := fresh "f" in
- let H := fresh in
- set (k := v);
- set (k' := @proj1_sig A B k);
- pose proof (proj2_sig k) as H;
- change (proj1_sig k) with k' in H;
- clearbody k'; clear k;
- cbv beta in *
- end.
- cbv [Interp Curry.curry2] in *.
- unfold interpf, interpf_step; fold_interpf.
- cbv [ladderstep_other_assoc interp_flat_type GF25519.fe25519].
- Time
- abstract (
- repeat match goal with
- | [ |- (dlet x := ?y in @?z x) = (dlet x' := ?y' in @?z' x') ]
- => refine ((fun pf0 pf1 => @Proper_Let_In_nd_changebody _ _ Logic.eq _ y y' pf0 z z' pf1)
- (_ : y = y')
- (_ : forall x, z x = z' x));
- cbv beta; intros;
- [ cbv [Let_In Common.ExprBinOpT] | ]
- end;
- repeat match goal with
- | _ => rewrite !interpf_invert_Abs
- | _ => rewrite_hyp !*
- | _ => progress cbv [interp_base_type]
- | [ |- ?x = ?x ] => reflexivity
- | _ => rewrite <- !surjective_pairing
- end
- ).
-Time Defined.
-Lemma rladderstepZ_sigP : rexpr_sigP _ uncurried_ladderstep rladderstepZ''.
-Proof.
- exact rladderstepZ_sigP'.
-Qed.
-Definition rladderstepZ_sig
- := exist (fun v => rexpr_sigP _ _ v) rladderstepZ'' rladderstepZ_sigP.
-
-Definition rladderstep_input_bounds
- := (ExprUnOp_bounds, ExprUnOp_bounds,
- ((ExprUnOp_bounds, ExprUnOp_bounds),
- (ExprUnOp_bounds, ExprUnOp_bounds))).
-
-Time Definition rladderstepW := Eval vm_compute in rword_of_Z rladderstepZ_sig.
-Lemma rladderstepW_correct_and_bounded_gen : correct_and_bounded_genT rladderstepW rladderstepZ_sig.
-Proof. Time rexpr_correct. Time Qed.
-Definition rladderstep_output_bounds := Eval vm_compute in compute_bounds rladderstepW rladderstep_input_bounds.
-
-Local Obligation Tactic := intros; vm_compute; constructor.
-
-(*
-Program Definition rladderstepW_correct_and_bounded
- := Expr9_4Op_correct_and_bounded
- rladderstepW uncurried_ladderstep rladderstepZ_sig rladderstepW_correct_and_bounded_gen
- _ _.
- *)
-
-Local Open Scope string_scope.
-Compute ("Ladderstep", compute_bounds_for_display rladderstepW rladderstep_input_bounds).
-Compute ("Ladderstep overflows? ", sanity_compute rladderstepW rladderstep_input_bounds).
-Compute ("Ladderstep overflows (error if it does)? ", sanity_check rladderstepW rladderstep_input_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/LadderStepDisplay.log b/src/Specific/GF25519Reflective/Reified/LadderStepDisplay.log
deleted file mode 100644
index ce961efd5..000000000
--- a/src/Specific/GF25519Reflective/Reified/LadderStepDisplay.log
+++ /dev/null
@@ -1,3459 +0,0 @@
-rladderstepW =
-fun var : base_type -> Type =>
-_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ x29 x30 x31 x32
- x33 x34 x35 x36 x37 x38 x39 x40 x41 x42 x43 x44 x45 x46 x47 x48 x49 x50 x51
- x52 x53 x54 x55 x56 x57 x58 x59 x60 x61 x62 x63 x64 x65 x66 x67 x68 x69 x70
- x71 x72 x73 x74 x75 x76 x77 x78 x79 x80 x81 x82 x83 x84 x85 x86 x87
- x88 : var TZ,
-Tbase TZ x89 = x49 + x59;
-Tbase TZ x90 = x50 + x60;
-Tbase TZ x91 = x51 + x61;
-Tbase TZ x92 = x52 + x62;
-Tbase TZ x93 = x53 + x63;
-Tbase TZ x94 = x54 + x64;
-Tbase TZ x95 = x55 + x65;
-Tbase TZ x96 = x56 + x66;
-Tbase TZ x97 = x57 + x67;
-Tbase TZ x98 = x58 + x68;
-Tbase TZ x99 = x98 * x98;
-Tbase TZ x100 = 0x2;
-Tbase TZ x101 = x97 * x100;
-Tbase TZ x102 = x89 * x101;
-Tbase TZ x103 = x90 * x96;
-Tbase TZ x104 = 0x2;
-Tbase TZ x105 = x95 * x104;
-Tbase TZ x106 = x91 * x105;
-Tbase TZ x107 = x92 * x94;
-Tbase TZ x108 = 0x2;
-Tbase TZ x109 = x93 * x108;
-Tbase TZ x110 = x93 * x109;
-Tbase TZ x111 = x94 * x92;
-Tbase TZ x112 = 0x2;
-Tbase TZ x113 = x91 * x112;
-Tbase TZ x114 = x95 * x113;
-Tbase TZ x115 = x96 * x90;
-Tbase TZ x116 = 0x2;
-Tbase TZ x117 = x89 * x116;
-Tbase TZ x118 = x97 * x117;
-Tbase TZ x119 = x115 + x118;
-Tbase TZ x120 = x114 + x119;
-Tbase TZ x121 = x111 + x120;
-Tbase TZ x122 = x110 + x121;
-Tbase TZ x123 = x107 + x122;
-Tbase TZ x124 = x106 + x123;
-Tbase TZ x125 = x103 + x124;
-Tbase TZ x126 = x102 + x125;
-Tbase TZ x127 = 0x13;
-Tbase TZ x128 = x127 * x126;
-Tbase TZ x129 = x99 + x128;
-Tbase TZ x130 = 0x1a;
-Tbase TZ x131 = x129 >> x130;
-Tbase TZ x132 = x97 * x98;
-Tbase TZ x133 = x98 * x97;
-Tbase TZ x134 = x132 + x133;
-Tbase TZ x135 = x89 * x96;
-Tbase TZ x136 = x90 * x95;
-Tbase TZ x137 = x91 * x94;
-Tbase TZ x138 = x92 * x93;
-Tbase TZ x139 = x93 * x92;
-Tbase TZ x140 = x94 * x91;
-Tbase TZ x141 = x95 * x90;
-Tbase TZ x142 = x96 * x89;
-Tbase TZ x143 = x141 + x142;
-Tbase TZ x144 = x140 + x143;
-Tbase TZ x145 = x139 + x144;
-Tbase TZ x146 = x138 + x145;
-Tbase TZ x147 = x137 + x146;
-Tbase TZ x148 = x136 + x147;
-Tbase TZ x149 = x135 + x148;
-Tbase TZ x150 = 0x13;
-Tbase TZ x151 = x150 * x149;
-Tbase TZ x152 = x134 + x151;
-Tbase TZ x153 = x131 + x152;
-Tbase TZ x154 = 0x19;
-Tbase TZ x155 = x153 >> x154;
-Tbase TZ x156 = x96 * x98;
-Tbase TZ x157 = 0x2;
-Tbase TZ x158 = x97 * x157;
-Tbase TZ x159 = x97 * x158;
-Tbase TZ x160 = x98 * x96;
-Tbase TZ x161 = x159 + x160;
-Tbase TZ x162 = x156 + x161;
-Tbase TZ x163 = 0x2;
-Tbase TZ x164 = x95 * x163;
-Tbase TZ x165 = x89 * x164;
-Tbase TZ x166 = x90 * x94;
-Tbase TZ x167 = 0x2;
-Tbase TZ x168 = x93 * x167;
-Tbase TZ x169 = x91 * x168;
-Tbase TZ x170 = x92 * x92;
-Tbase TZ x171 = 0x2;
-Tbase TZ x172 = x91 * x171;
-Tbase TZ x173 = x93 * x172;
-Tbase TZ x174 = x94 * x90;
-Tbase TZ x175 = 0x2;
-Tbase TZ x176 = x89 * x175;
-Tbase TZ x177 = x95 * x176;
-Tbase TZ x178 = x174 + x177;
-Tbase TZ x179 = x173 + x178;
-Tbase TZ x180 = x170 + x179;
-Tbase TZ x181 = x169 + x180;
-Tbase TZ x182 = x166 + x181;
-Tbase TZ x183 = x165 + x182;
-Tbase TZ x184 = 0x13;
-Tbase TZ x185 = x184 * x183;
-Tbase TZ x186 = x162 + x185;
-Tbase TZ x187 = x155 + x186;
-Tbase TZ x188 = 0x1a;
-Tbase TZ x189 = x187 >> x188;
-Tbase TZ x190 = x95 * x98;
-Tbase TZ x191 = x96 * x97;
-Tbase TZ x192 = x97 * x96;
-Tbase TZ x193 = x98 * x95;
-Tbase TZ x194 = x192 + x193;
-Tbase TZ x195 = x191 + x194;
-Tbase TZ x196 = x190 + x195;
-Tbase TZ x197 = x89 * x94;
-Tbase TZ x198 = x90 * x93;
-Tbase TZ x199 = x91 * x92;
-Tbase TZ x200 = x92 * x91;
-Tbase TZ x201 = x93 * x90;
-Tbase TZ x202 = x94 * x89;
-Tbase TZ x203 = x201 + x202;
-Tbase TZ x204 = x200 + x203;
-Tbase TZ x205 = x199 + x204;
-Tbase TZ x206 = x198 + x205;
-Tbase TZ x207 = x197 + x206;
-Tbase TZ x208 = 0x13;
-Tbase TZ x209 = x208 * x207;
-Tbase TZ x210 = x196 + x209;
-Tbase TZ x211 = x189 + x210;
-Tbase TZ x212 = 0x19;
-Tbase TZ x213 = x211 >> x212;
-Tbase TZ x214 = x94 * x98;
-Tbase TZ x215 = 0x2;
-Tbase TZ x216 = x97 * x215;
-Tbase TZ x217 = x95 * x216;
-Tbase TZ x218 = x96 * x96;
-Tbase TZ x219 = 0x2;
-Tbase TZ x220 = x95 * x219;
-Tbase TZ x221 = x97 * x220;
-Tbase TZ x222 = x98 * x94;
-Tbase TZ x223 = x221 + x222;
-Tbase TZ x224 = x218 + x223;
-Tbase TZ x225 = x217 + x224;
-Tbase TZ x226 = x214 + x225;
-Tbase TZ x227 = 0x2;
-Tbase TZ x228 = x93 * x227;
-Tbase TZ x229 = x89 * x228;
-Tbase TZ x230 = x90 * x92;
-Tbase TZ x231 = 0x2;
-Tbase TZ x232 = x91 * x231;
-Tbase TZ x233 = x91 * x232;
-Tbase TZ x234 = x92 * x90;
-Tbase TZ x235 = 0x2;
-Tbase TZ x236 = x89 * x235;
-Tbase TZ x237 = x93 * x236;
-Tbase TZ x238 = x234 + x237;
-Tbase TZ x239 = x233 + x238;
-Tbase TZ x240 = x230 + x239;
-Tbase TZ x241 = x229 + x240;
-Tbase TZ x242 = 0x13;
-Tbase TZ x243 = x242 * x241;
-Tbase TZ x244 = x226 + x243;
-Tbase TZ x245 = x213 + x244;
-Tbase TZ x246 = 0x1a;
-Tbase TZ x247 = x245 >> x246;
-Tbase TZ x248 = x93 * x98;
-Tbase TZ x249 = x94 * x97;
-Tbase TZ x250 = x95 * x96;
-Tbase TZ x251 = x96 * x95;
-Tbase TZ x252 = x97 * x94;
-Tbase TZ x253 = x98 * x93;
-Tbase TZ x254 = x252 + x253;
-Tbase TZ x255 = x251 + x254;
-Tbase TZ x256 = x250 + x255;
-Tbase TZ x257 = x249 + x256;
-Tbase TZ x258 = x248 + x257;
-Tbase TZ x259 = x89 * x92;
-Tbase TZ x260 = x90 * x91;
-Tbase TZ x261 = x91 * x90;
-Tbase TZ x262 = x92 * x89;
-Tbase TZ x263 = x261 + x262;
-Tbase TZ x264 = x260 + x263;
-Tbase TZ x265 = x259 + x264;
-Tbase TZ x266 = 0x13;
-Tbase TZ x267 = x266 * x265;
-Tbase TZ x268 = x258 + x267;
-Tbase TZ x269 = x247 + x268;
-Tbase TZ x270 = 0x19;
-Tbase TZ x271 = x269 >> x270;
-Tbase TZ x272 = x92 * x98;
-Tbase TZ x273 = 0x2;
-Tbase TZ x274 = x97 * x273;
-Tbase TZ x275 = x93 * x274;
-Tbase TZ x276 = x94 * x96;
-Tbase TZ x277 = 0x2;
-Tbase TZ x278 = x95 * x277;
-Tbase TZ x279 = x95 * x278;
-Tbase TZ x280 = x96 * x94;
-Tbase TZ x281 = 0x2;
-Tbase TZ x282 = x93 * x281;
-Tbase TZ x283 = x97 * x282;
-Tbase TZ x284 = x98 * x92;
-Tbase TZ x285 = x283 + x284;
-Tbase TZ x286 = x280 + x285;
-Tbase TZ x287 = x279 + x286;
-Tbase TZ x288 = x276 + x287;
-Tbase TZ x289 = x275 + x288;
-Tbase TZ x290 = x272 + x289;
-Tbase TZ x291 = 0x2;
-Tbase TZ x292 = x91 * x291;
-Tbase TZ x293 = x89 * x292;
-Tbase TZ x294 = x90 * x90;
-Tbase TZ x295 = 0x2;
-Tbase TZ x296 = x89 * x295;
-Tbase TZ x297 = x91 * x296;
-Tbase TZ x298 = x294 + x297;
-Tbase TZ x299 = x293 + x298;
-Tbase TZ x300 = 0x13;
-Tbase TZ x301 = x300 * x299;
-Tbase TZ x302 = x290 + x301;
-Tbase TZ x303 = x271 + x302;
-Tbase TZ x304 = 0x1a;
-Tbase TZ x305 = x303 >> x304;
-Tbase TZ x306 = x91 * x98;
-Tbase TZ x307 = x92 * x97;
-Tbase TZ x308 = x93 * x96;
-Tbase TZ x309 = x94 * x95;
-Tbase TZ x310 = x95 * x94;
-Tbase TZ x311 = x96 * x93;
-Tbase TZ x312 = x97 * x92;
-Tbase TZ x313 = x98 * x91;
-Tbase TZ x314 = x312 + x313;
-Tbase TZ x315 = x311 + x314;
-Tbase TZ x316 = x310 + x315;
-Tbase TZ x317 = x309 + x316;
-Tbase TZ x318 = x308 + x317;
-Tbase TZ x319 = x307 + x318;
-Tbase TZ x320 = x306 + x319;
-Tbase TZ x321 = x89 * x90;
-Tbase TZ x322 = x90 * x89;
-Tbase TZ x323 = x321 + x322;
-Tbase TZ x324 = 0x13;
-Tbase TZ x325 = x324 * x323;
-Tbase TZ x326 = x320 + x325;
-Tbase TZ x327 = x305 + x326;
-Tbase TZ x328 = 0x19;
-Tbase TZ x329 = x327 >> x328;
-Tbase TZ x330 = x90 * x98;
-Tbase TZ x331 = 0x2;
-Tbase TZ x332 = x97 * x331;
-Tbase TZ x333 = x91 * x332;
-Tbase TZ x334 = x92 * x96;
-Tbase TZ x335 = 0x2;
-Tbase TZ x336 = x95 * x335;
-Tbase TZ x337 = x93 * x336;
-Tbase TZ x338 = x94 * x94;
-Tbase TZ x339 = 0x2;
-Tbase TZ x340 = x93 * x339;
-Tbase TZ x341 = x95 * x340;
-Tbase TZ x342 = x96 * x92;
-Tbase TZ x343 = 0x2;
-Tbase TZ x344 = x91 * x343;
-Tbase TZ x345 = x97 * x344;
-Tbase TZ x346 = x98 * x90;
-Tbase TZ x347 = x345 + x346;
-Tbase TZ x348 = x342 + x347;
-Tbase TZ x349 = x341 + x348;
-Tbase TZ x350 = x338 + x349;
-Tbase TZ x351 = x337 + x350;
-Tbase TZ x352 = x334 + x351;
-Tbase TZ x353 = x333 + x352;
-Tbase TZ x354 = x330 + x353;
-Tbase TZ x355 = 0x2;
-Tbase TZ x356 = x89 * x355;
-Tbase TZ x357 = x89 * x356;
-Tbase TZ x358 = 0x13;
-Tbase TZ x359 = x358 * x357;
-Tbase TZ x360 = x354 + x359;
-Tbase TZ x361 = x329 + x360;
-Tbase TZ x362 = 0x1a;
-Tbase TZ x363 = x361 >> x362;
-Tbase TZ x364 = x89 * x98;
-Tbase TZ x365 = x90 * x97;
-Tbase TZ x366 = x91 * x96;
-Tbase TZ x367 = x92 * x95;
-Tbase TZ x368 = x93 * x94;
-Tbase TZ x369 = x94 * x93;
-Tbase TZ x370 = x95 * x92;
-Tbase TZ x371 = x96 * x91;
-Tbase TZ x372 = x97 * x90;
-Tbase TZ x373 = x98 * x89;
-Tbase TZ x374 = x372 + x373;
-Tbase TZ x375 = x371 + x374;
-Tbase TZ x376 = x370 + x375;
-Tbase TZ x377 = x369 + x376;
-Tbase TZ x378 = x368 + x377;
-Tbase TZ x379 = x367 + x378;
-Tbase TZ x380 = x366 + x379;
-Tbase TZ x381 = x365 + x380;
-Tbase TZ x382 = x364 + x381;
-Tbase TZ x383 = x363 + x382;
-Tbase TZ x384 = 0x19;
-Tbase TZ x385 = x383 >> x384;
-Tbase TZ x386 = 0x13;
-Tbase TZ x387 = x386 * x385;
-Tbase TZ x388 = 0x3ffffff;
-Tbase TZ x389 = x129 & x388;
-Tbase TZ x390 = x387 + x389;
-Tbase TZ x391 = 0x1a;
-Tbase TZ x392 = x390 >> x391;
-Tbase TZ x393 = 0x1ffffff;
-Tbase TZ x394 = x153 & x393;
-Tbase TZ x395 = x392 + x394;
-Tbase TZ x396 = 0x1ffffff;
-Tbase TZ x397 = x383 & x396;
-Tbase TZ x398 = 0x3ffffff;
-Tbase TZ x399 = x361 & x398;
-Tbase TZ x400 = 0x1ffffff;
-Tbase TZ x401 = x327 & x400;
-Tbase TZ x402 = 0x3ffffff;
-Tbase TZ x403 = x303 & x402;
-Tbase TZ x404 = 0x1ffffff;
-Tbase TZ x405 = x269 & x404;
-Tbase TZ x406 = 0x3ffffff;
-Tbase TZ x407 = x245 & x406;
-Tbase TZ x408 = 0x1ffffff;
-Tbase TZ x409 = x211 & x408;
-Tbase TZ x410 = 0x19;
-Tbase TZ x411 = x395 >> x410;
-Tbase TZ x412 = 0x3ffffff;
-Tbase TZ x413 = x187 & x412;
-Tbase TZ x414 = x411 + x413;
-Tbase TZ x415 = 0x1ffffff;
-Tbase TZ x416 = x395 & x415;
-Tbase TZ x417 = 0x3ffffff;
-Tbase TZ x418 = x390 & x417;
-Tbase TZ x419 = 0x3fffffe;
-Tbase TZ x420 = x419 + x49;
-Tbase TZ x421 = x420 - x59;
-Tbase TZ x422 = 0x7fffffe;
-Tbase TZ x423 = x422 + x50;
-Tbase TZ x424 = x423 - x60;
-Tbase TZ x425 = 0x3fffffe;
-Tbase TZ x426 = x425 + x51;
-Tbase TZ x427 = x426 - x61;
-Tbase TZ x428 = 0x7fffffe;
-Tbase TZ x429 = x428 + x52;
-Tbase TZ x430 = x429 - x62;
-Tbase TZ x431 = 0x3fffffe;
-Tbase TZ x432 = x431 + x53;
-Tbase TZ x433 = x432 - x63;
-Tbase TZ x434 = 0x7fffffe;
-Tbase TZ x435 = x434 + x54;
-Tbase TZ x436 = x435 - x64;
-Tbase TZ x437 = 0x3fffffe;
-Tbase TZ x438 = x437 + x55;
-Tbase TZ x439 = x438 - x65;
-Tbase TZ x440 = 0x7fffffe;
-Tbase TZ x441 = x440 + x56;
-Tbase TZ x442 = x441 - x66;
-Tbase TZ x443 = 0x3fffffe;
-Tbase TZ x444 = x443 + x57;
-Tbase TZ x445 = x444 - x67;
-Tbase TZ x446 = 0x7ffffda;
-Tbase TZ x447 = x446 + x58;
-Tbase TZ x448 = x447 - x68;
-Tbase TZ x449 = x448 * x448;
-Tbase TZ x450 = 0x2;
-Tbase TZ x451 = x445 * x450;
-Tbase TZ x452 = x421 * x451;
-Tbase TZ x453 = x424 * x442;
-Tbase TZ x454 = 0x2;
-Tbase TZ x455 = x439 * x454;
-Tbase TZ x456 = x427 * x455;
-Tbase TZ x457 = x430 * x436;
-Tbase TZ x458 = 0x2;
-Tbase TZ x459 = x433 * x458;
-Tbase TZ x460 = x433 * x459;
-Tbase TZ x461 = x436 * x430;
-Tbase TZ x462 = 0x2;
-Tbase TZ x463 = x427 * x462;
-Tbase TZ x464 = x439 * x463;
-Tbase TZ x465 = x442 * x424;
-Tbase TZ x466 = 0x2;
-Tbase TZ x467 = x421 * x466;
-Tbase TZ x468 = x445 * x467;
-Tbase TZ x469 = x465 + x468;
-Tbase TZ x470 = x464 + x469;
-Tbase TZ x471 = x461 + x470;
-Tbase TZ x472 = x460 + x471;
-Tbase TZ x473 = x457 + x472;
-Tbase TZ x474 = x456 + x473;
-Tbase TZ x475 = x453 + x474;
-Tbase TZ x476 = x452 + x475;
-Tbase TZ x477 = 0x13;
-Tbase TZ x478 = x477 * x476;
-Tbase TZ x479 = x449 + x478;
-Tbase TZ x480 = 0x1a;
-Tbase TZ x481 = x479 >> x480;
-Tbase TZ x482 = x445 * x448;
-Tbase TZ x483 = x448 * x445;
-Tbase TZ x484 = x482 + x483;
-Tbase TZ x485 = x421 * x442;
-Tbase TZ x486 = x424 * x439;
-Tbase TZ x487 = x427 * x436;
-Tbase TZ x488 = x430 * x433;
-Tbase TZ x489 = x433 * x430;
-Tbase TZ x490 = x436 * x427;
-Tbase TZ x491 = x439 * x424;
-Tbase TZ x492 = x442 * x421;
-Tbase TZ x493 = x491 + x492;
-Tbase TZ x494 = x490 + x493;
-Tbase TZ x495 = x489 + x494;
-Tbase TZ x496 = x488 + x495;
-Tbase TZ x497 = x487 + x496;
-Tbase TZ x498 = x486 + x497;
-Tbase TZ x499 = x485 + x498;
-Tbase TZ x500 = 0x13;
-Tbase TZ x501 = x500 * x499;
-Tbase TZ x502 = x484 + x501;
-Tbase TZ x503 = x481 + x502;
-Tbase TZ x504 = 0x19;
-Tbase TZ x505 = x503 >> x504;
-Tbase TZ x506 = x442 * x448;
-Tbase TZ x507 = 0x2;
-Tbase TZ x508 = x445 * x507;
-Tbase TZ x509 = x445 * x508;
-Tbase TZ x510 = x448 * x442;
-Tbase TZ x511 = x509 + x510;
-Tbase TZ x512 = x506 + x511;
-Tbase TZ x513 = 0x2;
-Tbase TZ x514 = x439 * x513;
-Tbase TZ x515 = x421 * x514;
-Tbase TZ x516 = x424 * x436;
-Tbase TZ x517 = 0x2;
-Tbase TZ x518 = x433 * x517;
-Tbase TZ x519 = x427 * x518;
-Tbase TZ x520 = x430 * x430;
-Tbase TZ x521 = 0x2;
-Tbase TZ x522 = x427 * x521;
-Tbase TZ x523 = x433 * x522;
-Tbase TZ x524 = x436 * x424;
-Tbase TZ x525 = 0x2;
-Tbase TZ x526 = x421 * x525;
-Tbase TZ x527 = x439 * x526;
-Tbase TZ x528 = x524 + x527;
-Tbase TZ x529 = x523 + x528;
-Tbase TZ x530 = x520 + x529;
-Tbase TZ x531 = x519 + x530;
-Tbase TZ x532 = x516 + x531;
-Tbase TZ x533 = x515 + x532;
-Tbase TZ x534 = 0x13;
-Tbase TZ x535 = x534 * x533;
-Tbase TZ x536 = x512 + x535;
-Tbase TZ x537 = x505 + x536;
-Tbase TZ x538 = 0x1a;
-Tbase TZ x539 = x537 >> x538;
-Tbase TZ x540 = x439 * x448;
-Tbase TZ x541 = x442 * x445;
-Tbase TZ x542 = x445 * x442;
-Tbase TZ x543 = x448 * x439;
-Tbase TZ x544 = x542 + x543;
-Tbase TZ x545 = x541 + x544;
-Tbase TZ x546 = x540 + x545;
-Tbase TZ x547 = x421 * x436;
-Tbase TZ x548 = x424 * x433;
-Tbase TZ x549 = x427 * x430;
-Tbase TZ x550 = x430 * x427;
-Tbase TZ x551 = x433 * x424;
-Tbase TZ x552 = x436 * x421;
-Tbase TZ x553 = x551 + x552;
-Tbase TZ x554 = x550 + x553;
-Tbase TZ x555 = x549 + x554;
-Tbase TZ x556 = x548 + x555;
-Tbase TZ x557 = x547 + x556;
-Tbase TZ x558 = 0x13;
-Tbase TZ x559 = x558 * x557;
-Tbase TZ x560 = x546 + x559;
-Tbase TZ x561 = x539 + x560;
-Tbase TZ x562 = 0x19;
-Tbase TZ x563 = x561 >> x562;
-Tbase TZ x564 = x436 * x448;
-Tbase TZ x565 = 0x2;
-Tbase TZ x566 = x445 * x565;
-Tbase TZ x567 = x439 * x566;
-Tbase TZ x568 = x442 * x442;
-Tbase TZ x569 = 0x2;
-Tbase TZ x570 = x439 * x569;
-Tbase TZ x571 = x445 * x570;
-Tbase TZ x572 = x448 * x436;
-Tbase TZ x573 = x571 + x572;
-Tbase TZ x574 = x568 + x573;
-Tbase TZ x575 = x567 + x574;
-Tbase TZ x576 = x564 + x575;
-Tbase TZ x577 = 0x2;
-Tbase TZ x578 = x433 * x577;
-Tbase TZ x579 = x421 * x578;
-Tbase TZ x580 = x424 * x430;
-Tbase TZ x581 = 0x2;
-Tbase TZ x582 = x427 * x581;
-Tbase TZ x583 = x427 * x582;
-Tbase TZ x584 = x430 * x424;
-Tbase TZ x585 = 0x2;
-Tbase TZ x586 = x421 * x585;
-Tbase TZ x587 = x433 * x586;
-Tbase TZ x588 = x584 + x587;
-Tbase TZ x589 = x583 + x588;
-Tbase TZ x590 = x580 + x589;
-Tbase TZ x591 = x579 + x590;
-Tbase TZ x592 = 0x13;
-Tbase TZ x593 = x592 * x591;
-Tbase TZ x594 = x576 + x593;
-Tbase TZ x595 = x563 + x594;
-Tbase TZ x596 = 0x1a;
-Tbase TZ x597 = x595 >> x596;
-Tbase TZ x598 = x433 * x448;
-Tbase TZ x599 = x436 * x445;
-Tbase TZ x600 = x439 * x442;
-Tbase TZ x601 = x442 * x439;
-Tbase TZ x602 = x445 * x436;
-Tbase TZ x603 = x448 * x433;
-Tbase TZ x604 = x602 + x603;
-Tbase TZ x605 = x601 + x604;
-Tbase TZ x606 = x600 + x605;
-Tbase TZ x607 = x599 + x606;
-Tbase TZ x608 = x598 + x607;
-Tbase TZ x609 = x421 * x430;
-Tbase TZ x610 = x424 * x427;
-Tbase TZ x611 = x427 * x424;
-Tbase TZ x612 = x430 * x421;
-Tbase TZ x613 = x611 + x612;
-Tbase TZ x614 = x610 + x613;
-Tbase TZ x615 = x609 + x614;
-Tbase TZ x616 = 0x13;
-Tbase TZ x617 = x616 * x615;
-Tbase TZ x618 = x608 + x617;
-Tbase TZ x619 = x597 + x618;
-Tbase TZ x620 = 0x19;
-Tbase TZ x621 = x619 >> x620;
-Tbase TZ x622 = x430 * x448;
-Tbase TZ x623 = 0x2;
-Tbase TZ x624 = x445 * x623;
-Tbase TZ x625 = x433 * x624;
-Tbase TZ x626 = x436 * x442;
-Tbase TZ x627 = 0x2;
-Tbase TZ x628 = x439 * x627;
-Tbase TZ x629 = x439 * x628;
-Tbase TZ x630 = x442 * x436;
-Tbase TZ x631 = 0x2;
-Tbase TZ x632 = x433 * x631;
-Tbase TZ x633 = x445 * x632;
-Tbase TZ x634 = x448 * x430;
-Tbase TZ x635 = x633 + x634;
-Tbase TZ x636 = x630 + x635;
-Tbase TZ x637 = x629 + x636;
-Tbase TZ x638 = x626 + x637;
-Tbase TZ x639 = x625 + x638;
-Tbase TZ x640 = x622 + x639;
-Tbase TZ x641 = 0x2;
-Tbase TZ x642 = x427 * x641;
-Tbase TZ x643 = x421 * x642;
-Tbase TZ x644 = x424 * x424;
-Tbase TZ x645 = 0x2;
-Tbase TZ x646 = x421 * x645;
-Tbase TZ x647 = x427 * x646;
-Tbase TZ x648 = x644 + x647;
-Tbase TZ x649 = x643 + x648;
-Tbase TZ x650 = 0x13;
-Tbase TZ x651 = x650 * x649;
-Tbase TZ x652 = x640 + x651;
-Tbase TZ x653 = x621 + x652;
-Tbase TZ x654 = 0x1a;
-Tbase TZ x655 = x653 >> x654;
-Tbase TZ x656 = x427 * x448;
-Tbase TZ x657 = x430 * x445;
-Tbase TZ x658 = x433 * x442;
-Tbase TZ x659 = x436 * x439;
-Tbase TZ x660 = x439 * x436;
-Tbase TZ x661 = x442 * x433;
-Tbase TZ x662 = x445 * x430;
-Tbase TZ x663 = x448 * x427;
-Tbase TZ x664 = x662 + x663;
-Tbase TZ x665 = x661 + x664;
-Tbase TZ x666 = x660 + x665;
-Tbase TZ x667 = x659 + x666;
-Tbase TZ x668 = x658 + x667;
-Tbase TZ x669 = x657 + x668;
-Tbase TZ x670 = x656 + x669;
-Tbase TZ x671 = x421 * x424;
-Tbase TZ x672 = x424 * x421;
-Tbase TZ x673 = x671 + x672;
-Tbase TZ x674 = 0x13;
-Tbase TZ x675 = x674 * x673;
-Tbase TZ x676 = x670 + x675;
-Tbase TZ x677 = x655 + x676;
-Tbase TZ x678 = 0x19;
-Tbase TZ x679 = x677 >> x678;
-Tbase TZ x680 = x424 * x448;
-Tbase TZ x681 = 0x2;
-Tbase TZ x682 = x445 * x681;
-Tbase TZ x683 = x427 * x682;
-Tbase TZ x684 = x430 * x442;
-Tbase TZ x685 = 0x2;
-Tbase TZ x686 = x439 * x685;
-Tbase TZ x687 = x433 * x686;
-Tbase TZ x688 = x436 * x436;
-Tbase TZ x689 = 0x2;
-Tbase TZ x690 = x433 * x689;
-Tbase TZ x691 = x439 * x690;
-Tbase TZ x692 = x442 * x430;
-Tbase TZ x693 = 0x2;
-Tbase TZ x694 = x427 * x693;
-Tbase TZ x695 = x445 * x694;
-Tbase TZ x696 = x448 * x424;
-Tbase TZ x697 = x695 + x696;
-Tbase TZ x698 = x692 + x697;
-Tbase TZ x699 = x691 + x698;
-Tbase TZ x700 = x688 + x699;
-Tbase TZ x701 = x687 + x700;
-Tbase TZ x702 = x684 + x701;
-Tbase TZ x703 = x683 + x702;
-Tbase TZ x704 = x680 + x703;
-Tbase TZ x705 = 0x2;
-Tbase TZ x706 = x421 * x705;
-Tbase TZ x707 = x421 * x706;
-Tbase TZ x708 = 0x13;
-Tbase TZ x709 = x708 * x707;
-Tbase TZ x710 = x704 + x709;
-Tbase TZ x711 = x679 + x710;
-Tbase TZ x712 = 0x1a;
-Tbase TZ x713 = x711 >> x712;
-Tbase TZ x714 = x421 * x448;
-Tbase TZ x715 = x424 * x445;
-Tbase TZ x716 = x427 * x442;
-Tbase TZ x717 = x430 * x439;
-Tbase TZ x718 = x433 * x436;
-Tbase TZ x719 = x436 * x433;
-Tbase TZ x720 = x439 * x430;
-Tbase TZ x721 = x442 * x427;
-Tbase TZ x722 = x445 * x424;
-Tbase TZ x723 = x448 * x421;
-Tbase TZ x724 = x722 + x723;
-Tbase TZ x725 = x721 + x724;
-Tbase TZ x726 = x720 + x725;
-Tbase TZ x727 = x719 + x726;
-Tbase TZ x728 = x718 + x727;
-Tbase TZ x729 = x717 + x728;
-Tbase TZ x730 = x716 + x729;
-Tbase TZ x731 = x715 + x730;
-Tbase TZ x732 = x714 + x731;
-Tbase TZ x733 = x713 + x732;
-Tbase TZ x734 = 0x19;
-Tbase TZ x735 = x733 >> x734;
-Tbase TZ x736 = 0x13;
-Tbase TZ x737 = x736 * x735;
-Tbase TZ x738 = 0x3ffffff;
-Tbase TZ x739 = x479 & x738;
-Tbase TZ x740 = x737 + x739;
-Tbase TZ x741 = 0x1a;
-Tbase TZ x742 = x740 >> x741;
-Tbase TZ x743 = 0x1ffffff;
-Tbase TZ x744 = x503 & x743;
-Tbase TZ x745 = x742 + x744;
-Tbase TZ x746 = 0x1ffffff;
-Tbase TZ x747 = x733 & x746;
-Tbase TZ x748 = 0x3ffffff;
-Tbase TZ x749 = x711 & x748;
-Tbase TZ x750 = 0x1ffffff;
-Tbase TZ x751 = x677 & x750;
-Tbase TZ x752 = 0x3ffffff;
-Tbase TZ x753 = x653 & x752;
-Tbase TZ x754 = 0x1ffffff;
-Tbase TZ x755 = x619 & x754;
-Tbase TZ x756 = 0x3ffffff;
-Tbase TZ x757 = x595 & x756;
-Tbase TZ x758 = 0x1ffffff;
-Tbase TZ x759 = x561 & x758;
-Tbase TZ x760 = 0x19;
-Tbase TZ x761 = x745 >> x760;
-Tbase TZ x762 = 0x3ffffff;
-Tbase TZ x763 = x537 & x762;
-Tbase TZ x764 = x761 + x763;
-Tbase TZ x765 = 0x1ffffff;
-Tbase TZ x766 = x745 & x765;
-Tbase TZ x767 = 0x3ffffff;
-Tbase TZ x768 = x740 & x767;
-Tbase TZ x769 = 0x3fffffe;
-Tbase TZ x770 = x769 + x397;
-Tbase TZ x771 = x770 - x747;
-Tbase TZ x772 = 0x7fffffe;
-Tbase TZ x773 = x772 + x399;
-Tbase TZ x774 = x773 - x749;
-Tbase TZ x775 = 0x3fffffe;
-Tbase TZ x776 = x775 + x401;
-Tbase TZ x777 = x776 - x751;
-Tbase TZ x778 = 0x7fffffe;
-Tbase TZ x779 = x778 + x403;
-Tbase TZ x780 = x779 - x753;
-Tbase TZ x781 = 0x3fffffe;
-Tbase TZ x782 = x781 + x405;
-Tbase TZ x783 = x782 - x755;
-Tbase TZ x784 = 0x7fffffe;
-Tbase TZ x785 = x784 + x407;
-Tbase TZ x786 = x785 - x757;
-Tbase TZ x787 = 0x3fffffe;
-Tbase TZ x788 = x787 + x409;
-Tbase TZ x789 = x788 - x759;
-Tbase TZ x790 = 0x7fffffe;
-Tbase TZ x791 = x790 + x414;
-Tbase TZ x792 = x791 - x764;
-Tbase TZ x793 = 0x3fffffe;
-Tbase TZ x794 = x793 + x416;
-Tbase TZ x795 = x794 - x766;
-Tbase TZ x796 = 0x7ffffda;
-Tbase TZ x797 = x796 + x418;
-Tbase TZ x798 = x797 - x768;
-Tbase TZ x799 = x69 + x79;
-Tbase TZ x800 = x70 + x80;
-Tbase TZ x801 = x71 + x81;
-Tbase TZ x802 = x72 + x82;
-Tbase TZ x803 = x73 + x83;
-Tbase TZ x804 = x74 + x84;
-Tbase TZ x805 = x75 + x85;
-Tbase TZ x806 = x76 + x86;
-Tbase TZ x807 = x77 + x87;
-Tbase TZ x808 = x78 + x88;
-Tbase TZ x809 = 0x3fffffe;
-Tbase TZ x810 = x809 + x69;
-Tbase TZ x811 = x810 - x79;
-Tbase TZ x812 = 0x7fffffe;
-Tbase TZ x813 = x812 + x70;
-Tbase TZ x814 = x813 - x80;
-Tbase TZ x815 = 0x3fffffe;
-Tbase TZ x816 = x815 + x71;
-Tbase TZ x817 = x816 - x81;
-Tbase TZ x818 = 0x7fffffe;
-Tbase TZ x819 = x818 + x72;
-Tbase TZ x820 = x819 - x82;
-Tbase TZ x821 = 0x3fffffe;
-Tbase TZ x822 = x821 + x73;
-Tbase TZ x823 = x822 - x83;
-Tbase TZ x824 = 0x7fffffe;
-Tbase TZ x825 = x824 + x74;
-Tbase TZ x826 = x825 - x84;
-Tbase TZ x827 = 0x3fffffe;
-Tbase TZ x828 = x827 + x75;
-Tbase TZ x829 = x828 - x85;
-Tbase TZ x830 = 0x7fffffe;
-Tbase TZ x831 = x830 + x76;
-Tbase TZ x832 = x831 - x86;
-Tbase TZ x833 = 0x3fffffe;
-Tbase TZ x834 = x833 + x77;
-Tbase TZ x835 = x834 - x87;
-Tbase TZ x836 = 0x7ffffda;
-Tbase TZ x837 = x836 + x78;
-Tbase TZ x838 = x837 - x88;
-Tbase TZ x839 = x838 * x98;
-Tbase TZ x840 = 0x2;
-Tbase TZ x841 = x97 * x840;
-Tbase TZ x842 = x811 * x841;
-Tbase TZ x843 = x814 * x96;
-Tbase TZ x844 = 0x2;
-Tbase TZ x845 = x95 * x844;
-Tbase TZ x846 = x817 * x845;
-Tbase TZ x847 = x820 * x94;
-Tbase TZ x848 = 0x2;
-Tbase TZ x849 = x93 * x848;
-Tbase TZ x850 = x823 * x849;
-Tbase TZ x851 = x826 * x92;
-Tbase TZ x852 = 0x2;
-Tbase TZ x853 = x91 * x852;
-Tbase TZ x854 = x829 * x853;
-Tbase TZ x855 = x832 * x90;
-Tbase TZ x856 = 0x2;
-Tbase TZ x857 = x89 * x856;
-Tbase TZ x858 = x835 * x857;
-Tbase TZ x859 = x855 + x858;
-Tbase TZ x860 = x854 + x859;
-Tbase TZ x861 = x851 + x860;
-Tbase TZ x862 = x850 + x861;
-Tbase TZ x863 = x847 + x862;
-Tbase TZ x864 = x846 + x863;
-Tbase TZ x865 = x843 + x864;
-Tbase TZ x866 = x842 + x865;
-Tbase TZ x867 = 0x13;
-Tbase TZ x868 = x867 * x866;
-Tbase TZ x869 = x839 + x868;
-Tbase TZ x870 = 0x1a;
-Tbase TZ x871 = x869 >> x870;
-Tbase TZ x872 = x835 * x98;
-Tbase TZ x873 = x838 * x97;
-Tbase TZ x874 = x872 + x873;
-Tbase TZ x875 = x811 * x96;
-Tbase TZ x876 = x814 * x95;
-Tbase TZ x877 = x817 * x94;
-Tbase TZ x878 = x820 * x93;
-Tbase TZ x879 = x823 * x92;
-Tbase TZ x880 = x826 * x91;
-Tbase TZ x881 = x829 * x90;
-Tbase TZ x882 = x832 * x89;
-Tbase TZ x883 = x881 + x882;
-Tbase TZ x884 = x880 + x883;
-Tbase TZ x885 = x879 + x884;
-Tbase TZ x886 = x878 + x885;
-Tbase TZ x887 = x877 + x886;
-Tbase TZ x888 = x876 + x887;
-Tbase TZ x889 = x875 + x888;
-Tbase TZ x890 = 0x13;
-Tbase TZ x891 = x890 * x889;
-Tbase TZ x892 = x874 + x891;
-Tbase TZ x893 = x871 + x892;
-Tbase TZ x894 = 0x19;
-Tbase TZ x895 = x893 >> x894;
-Tbase TZ x896 = x832 * x98;
-Tbase TZ x897 = 0x2;
-Tbase TZ x898 = x97 * x897;
-Tbase TZ x899 = x835 * x898;
-Tbase TZ x900 = x838 * x96;
-Tbase TZ x901 = x899 + x900;
-Tbase TZ x902 = x896 + x901;
-Tbase TZ x903 = 0x2;
-Tbase TZ x904 = x95 * x903;
-Tbase TZ x905 = x811 * x904;
-Tbase TZ x906 = x814 * x94;
-Tbase TZ x907 = 0x2;
-Tbase TZ x908 = x93 * x907;
-Tbase TZ x909 = x817 * x908;
-Tbase TZ x910 = x820 * x92;
-Tbase TZ x911 = 0x2;
-Tbase TZ x912 = x91 * x911;
-Tbase TZ x913 = x823 * x912;
-Tbase TZ x914 = x826 * x90;
-Tbase TZ x915 = 0x2;
-Tbase TZ x916 = x89 * x915;
-Tbase TZ x917 = x829 * x916;
-Tbase TZ x918 = x914 + x917;
-Tbase TZ x919 = x913 + x918;
-Tbase TZ x920 = x910 + x919;
-Tbase TZ x921 = x909 + x920;
-Tbase TZ x922 = x906 + x921;
-Tbase TZ x923 = x905 + x922;
-Tbase TZ x924 = 0x13;
-Tbase TZ x925 = x924 * x923;
-Tbase TZ x926 = x902 + x925;
-Tbase TZ x927 = x895 + x926;
-Tbase TZ x928 = 0x1a;
-Tbase TZ x929 = x927 >> x928;
-Tbase TZ x930 = x829 * x98;
-Tbase TZ x931 = x832 * x97;
-Tbase TZ x932 = x835 * x96;
-Tbase TZ x933 = x838 * x95;
-Tbase TZ x934 = x932 + x933;
-Tbase TZ x935 = x931 + x934;
-Tbase TZ x936 = x930 + x935;
-Tbase TZ x937 = x811 * x94;
-Tbase TZ x938 = x814 * x93;
-Tbase TZ x939 = x817 * x92;
-Tbase TZ x940 = x820 * x91;
-Tbase TZ x941 = x823 * x90;
-Tbase TZ x942 = x826 * x89;
-Tbase TZ x943 = x941 + x942;
-Tbase TZ x944 = x940 + x943;
-Tbase TZ x945 = x939 + x944;
-Tbase TZ x946 = x938 + x945;
-Tbase TZ x947 = x937 + x946;
-Tbase TZ x948 = 0x13;
-Tbase TZ x949 = x948 * x947;
-Tbase TZ x950 = x936 + x949;
-Tbase TZ x951 = x929 + x950;
-Tbase TZ x952 = 0x19;
-Tbase TZ x953 = x951 >> x952;
-Tbase TZ x954 = x826 * x98;
-Tbase TZ x955 = 0x2;
-Tbase TZ x956 = x97 * x955;
-Tbase TZ x957 = x829 * x956;
-Tbase TZ x958 = x832 * x96;
-Tbase TZ x959 = 0x2;
-Tbase TZ x960 = x95 * x959;
-Tbase TZ x961 = x835 * x960;
-Tbase TZ x962 = x838 * x94;
-Tbase TZ x963 = x961 + x962;
-Tbase TZ x964 = x958 + x963;
-Tbase TZ x965 = x957 + x964;
-Tbase TZ x966 = x954 + x965;
-Tbase TZ x967 = 0x2;
-Tbase TZ x968 = x93 * x967;
-Tbase TZ x969 = x811 * x968;
-Tbase TZ x970 = x814 * x92;
-Tbase TZ x971 = 0x2;
-Tbase TZ x972 = x91 * x971;
-Tbase TZ x973 = x817 * x972;
-Tbase TZ x974 = x820 * x90;
-Tbase TZ x975 = 0x2;
-Tbase TZ x976 = x89 * x975;
-Tbase TZ x977 = x823 * x976;
-Tbase TZ x978 = x974 + x977;
-Tbase TZ x979 = x973 + x978;
-Tbase TZ x980 = x970 + x979;
-Tbase TZ x981 = x969 + x980;
-Tbase TZ x982 = 0x13;
-Tbase TZ x983 = x982 * x981;
-Tbase TZ x984 = x966 + x983;
-Tbase TZ x985 = x953 + x984;
-Tbase TZ x986 = 0x1a;
-Tbase TZ x987 = x985 >> x986;
-Tbase TZ x988 = x823 * x98;
-Tbase TZ x989 = x826 * x97;
-Tbase TZ x990 = x829 * x96;
-Tbase TZ x991 = x832 * x95;
-Tbase TZ x992 = x835 * x94;
-Tbase TZ x993 = x838 * x93;
-Tbase TZ x994 = x992 + x993;
-Tbase TZ x995 = x991 + x994;
-Tbase TZ x996 = x990 + x995;
-Tbase TZ x997 = x989 + x996;
-Tbase TZ x998 = x988 + x997;
-Tbase TZ x999 = x811 * x92;
-Tbase TZ x1000 = x814 * x91;
-Tbase TZ x1001 = x817 * x90;
-Tbase TZ x1002 = x820 * x89;
-Tbase TZ x1003 = x1001 + x1002;
-Tbase TZ x1004 = x1000 + x1003;
-Tbase TZ x1005 = x999 + x1004;
-Tbase TZ x1006 = 0x13;
-Tbase TZ x1007 = x1006 * x1005;
-Tbase TZ x1008 = x998 + x1007;
-Tbase TZ x1009 = x987 + x1008;
-Tbase TZ x1010 = 0x19;
-Tbase TZ x1011 = x1009 >> x1010;
-Tbase TZ x1012 = x820 * x98;
-Tbase TZ x1013 = 0x2;
-Tbase TZ x1014 = x97 * x1013;
-Tbase TZ x1015 = x823 * x1014;
-Tbase TZ x1016 = x826 * x96;
-Tbase TZ x1017 = 0x2;
-Tbase TZ x1018 = x95 * x1017;
-Tbase TZ x1019 = x829 * x1018;
-Tbase TZ x1020 = x832 * x94;
-Tbase TZ x1021 = 0x2;
-Tbase TZ x1022 = x93 * x1021;
-Tbase TZ x1023 = x835 * x1022;
-Tbase TZ x1024 = x838 * x92;
-Tbase TZ x1025 = x1023 + x1024;
-Tbase TZ x1026 = x1020 + x1025;
-Tbase TZ x1027 = x1019 + x1026;
-Tbase TZ x1028 = x1016 + x1027;
-Tbase TZ x1029 = x1015 + x1028;
-Tbase TZ x1030 = x1012 + x1029;
-Tbase TZ x1031 = 0x2;
-Tbase TZ x1032 = x91 * x1031;
-Tbase TZ x1033 = x811 * x1032;
-Tbase TZ x1034 = x814 * x90;
-Tbase TZ x1035 = 0x2;
-Tbase TZ x1036 = x89 * x1035;
-Tbase TZ x1037 = x817 * x1036;
-Tbase TZ x1038 = x1034 + x1037;
-Tbase TZ x1039 = x1033 + x1038;
-Tbase TZ x1040 = 0x13;
-Tbase TZ x1041 = x1040 * x1039;
-Tbase TZ x1042 = x1030 + x1041;
-Tbase TZ x1043 = x1011 + x1042;
-Tbase TZ x1044 = 0x1a;
-Tbase TZ x1045 = x1043 >> x1044;
-Tbase TZ x1046 = x817 * x98;
-Tbase TZ x1047 = x820 * x97;
-Tbase TZ x1048 = x823 * x96;
-Tbase TZ x1049 = x826 * x95;
-Tbase TZ x1050 = x829 * x94;
-Tbase TZ x1051 = x832 * x93;
-Tbase TZ x1052 = x835 * x92;
-Tbase TZ x1053 = x838 * x91;
-Tbase TZ x1054 = x1052 + x1053;
-Tbase TZ x1055 = x1051 + x1054;
-Tbase TZ x1056 = x1050 + x1055;
-Tbase TZ x1057 = x1049 + x1056;
-Tbase TZ x1058 = x1048 + x1057;
-Tbase TZ x1059 = x1047 + x1058;
-Tbase TZ x1060 = x1046 + x1059;
-Tbase TZ x1061 = x811 * x90;
-Tbase TZ x1062 = x814 * x89;
-Tbase TZ x1063 = x1061 + x1062;
-Tbase TZ x1064 = 0x13;
-Tbase TZ x1065 = x1064 * x1063;
-Tbase TZ x1066 = x1060 + x1065;
-Tbase TZ x1067 = x1045 + x1066;
-Tbase TZ x1068 = 0x19;
-Tbase TZ x1069 = x1067 >> x1068;
-Tbase TZ x1070 = x814 * x98;
-Tbase TZ x1071 = 0x2;
-Tbase TZ x1072 = x97 * x1071;
-Tbase TZ x1073 = x817 * x1072;
-Tbase TZ x1074 = x820 * x96;
-Tbase TZ x1075 = 0x2;
-Tbase TZ x1076 = x95 * x1075;
-Tbase TZ x1077 = x823 * x1076;
-Tbase TZ x1078 = x826 * x94;
-Tbase TZ x1079 = 0x2;
-Tbase TZ x1080 = x93 * x1079;
-Tbase TZ x1081 = x829 * x1080;
-Tbase TZ x1082 = x832 * x92;
-Tbase TZ x1083 = 0x2;
-Tbase TZ x1084 = x91 * x1083;
-Tbase TZ x1085 = x835 * x1084;
-Tbase TZ x1086 = x838 * x90;
-Tbase TZ x1087 = x1085 + x1086;
-Tbase TZ x1088 = x1082 + x1087;
-Tbase TZ x1089 = x1081 + x1088;
-Tbase TZ x1090 = x1078 + x1089;
-Tbase TZ x1091 = x1077 + x1090;
-Tbase TZ x1092 = x1074 + x1091;
-Tbase TZ x1093 = x1073 + x1092;
-Tbase TZ x1094 = x1070 + x1093;
-Tbase TZ x1095 = 0x2;
-Tbase TZ x1096 = x89 * x1095;
-Tbase TZ x1097 = x811 * x1096;
-Tbase TZ x1098 = 0x13;
-Tbase TZ x1099 = x1098 * x1097;
-Tbase TZ x1100 = x1094 + x1099;
-Tbase TZ x1101 = x1069 + x1100;
-Tbase TZ x1102 = 0x1a;
-Tbase TZ x1103 = x1101 >> x1102;
-Tbase TZ x1104 = x811 * x98;
-Tbase TZ x1105 = x814 * x97;
-Tbase TZ x1106 = x817 * x96;
-Tbase TZ x1107 = x820 * x95;
-Tbase TZ x1108 = x823 * x94;
-Tbase TZ x1109 = x826 * x93;
-Tbase TZ x1110 = x829 * x92;
-Tbase TZ x1111 = x832 * x91;
-Tbase TZ x1112 = x835 * x90;
-Tbase TZ x1113 = x838 * x89;
-Tbase TZ x1114 = x1112 + x1113;
-Tbase TZ x1115 = x1111 + x1114;
-Tbase TZ x1116 = x1110 + x1115;
-Tbase TZ x1117 = x1109 + x1116;
-Tbase TZ x1118 = x1108 + x1117;
-Tbase TZ x1119 = x1107 + x1118;
-Tbase TZ x1120 = x1106 + x1119;
-Tbase TZ x1121 = x1105 + x1120;
-Tbase TZ x1122 = x1104 + x1121;
-Tbase TZ x1123 = x1103 + x1122;
-Tbase TZ x1124 = 0x19;
-Tbase TZ x1125 = x1123 >> x1124;
-Tbase TZ x1126 = 0x13;
-Tbase TZ x1127 = x1126 * x1125;
-Tbase TZ x1128 = 0x3ffffff;
-Tbase TZ x1129 = x869 & x1128;
-Tbase TZ x1130 = x1127 + x1129;
-Tbase TZ x1131 = 0x1a;
-Tbase TZ x1132 = x1130 >> x1131;
-Tbase TZ x1133 = 0x1ffffff;
-Tbase TZ x1134 = x893 & x1133;
-Tbase TZ x1135 = x1132 + x1134;
-Tbase TZ x1136 = 0x1ffffff;
-Tbase TZ x1137 = x1123 & x1136;
-Tbase TZ x1138 = 0x3ffffff;
-Tbase TZ x1139 = x1101 & x1138;
-Tbase TZ x1140 = 0x1ffffff;
-Tbase TZ x1141 = x1067 & x1140;
-Tbase TZ x1142 = 0x3ffffff;
-Tbase TZ x1143 = x1043 & x1142;
-Tbase TZ x1144 = 0x1ffffff;
-Tbase TZ x1145 = x1009 & x1144;
-Tbase TZ x1146 = 0x3ffffff;
-Tbase TZ x1147 = x985 & x1146;
-Tbase TZ x1148 = 0x1ffffff;
-Tbase TZ x1149 = x951 & x1148;
-Tbase TZ x1150 = 0x19;
-Tbase TZ x1151 = x1135 >> x1150;
-Tbase TZ x1152 = 0x3ffffff;
-Tbase TZ x1153 = x927 & x1152;
-Tbase TZ x1154 = x1151 + x1153;
-Tbase TZ x1155 = 0x1ffffff;
-Tbase TZ x1156 = x1135 & x1155;
-Tbase TZ x1157 = 0x3ffffff;
-Tbase TZ x1158 = x1130 & x1157;
-Tbase TZ x1159 = x808 * x448;
-Tbase TZ x1160 = 0x2;
-Tbase TZ x1161 = x445 * x1160;
-Tbase TZ x1162 = x799 * x1161;
-Tbase TZ x1163 = x800 * x442;
-Tbase TZ x1164 = 0x2;
-Tbase TZ x1165 = x439 * x1164;
-Tbase TZ x1166 = x801 * x1165;
-Tbase TZ x1167 = x802 * x436;
-Tbase TZ x1168 = 0x2;
-Tbase TZ x1169 = x433 * x1168;
-Tbase TZ x1170 = x803 * x1169;
-Tbase TZ x1171 = x804 * x430;
-Tbase TZ x1172 = 0x2;
-Tbase TZ x1173 = x427 * x1172;
-Tbase TZ x1174 = x805 * x1173;
-Tbase TZ x1175 = x806 * x424;
-Tbase TZ x1176 = 0x2;
-Tbase TZ x1177 = x421 * x1176;
-Tbase TZ x1178 = x807 * x1177;
-Tbase TZ x1179 = x1175 + x1178;
-Tbase TZ x1180 = x1174 + x1179;
-Tbase TZ x1181 = x1171 + x1180;
-Tbase TZ x1182 = x1170 + x1181;
-Tbase TZ x1183 = x1167 + x1182;
-Tbase TZ x1184 = x1166 + x1183;
-Tbase TZ x1185 = x1163 + x1184;
-Tbase TZ x1186 = x1162 + x1185;
-Tbase TZ x1187 = 0x13;
-Tbase TZ x1188 = x1187 * x1186;
-Tbase TZ x1189 = x1159 + x1188;
-Tbase TZ x1190 = 0x1a;
-Tbase TZ x1191 = x1189 >> x1190;
-Tbase TZ x1192 = x807 * x448;
-Tbase TZ x1193 = x808 * x445;
-Tbase TZ x1194 = x1192 + x1193;
-Tbase TZ x1195 = x799 * x442;
-Tbase TZ x1196 = x800 * x439;
-Tbase TZ x1197 = x801 * x436;
-Tbase TZ x1198 = x802 * x433;
-Tbase TZ x1199 = x803 * x430;
-Tbase TZ x1200 = x804 * x427;
-Tbase TZ x1201 = x805 * x424;
-Tbase TZ x1202 = x806 * x421;
-Tbase TZ x1203 = x1201 + x1202;
-Tbase TZ x1204 = x1200 + x1203;
-Tbase TZ x1205 = x1199 + x1204;
-Tbase TZ x1206 = x1198 + x1205;
-Tbase TZ x1207 = x1197 + x1206;
-Tbase TZ x1208 = x1196 + x1207;
-Tbase TZ x1209 = x1195 + x1208;
-Tbase TZ x1210 = 0x13;
-Tbase TZ x1211 = x1210 * x1209;
-Tbase TZ x1212 = x1194 + x1211;
-Tbase TZ x1213 = x1191 + x1212;
-Tbase TZ x1214 = 0x19;
-Tbase TZ x1215 = x1213 >> x1214;
-Tbase TZ x1216 = x806 * x448;
-Tbase TZ x1217 = 0x2;
-Tbase TZ x1218 = x445 * x1217;
-Tbase TZ x1219 = x807 * x1218;
-Tbase TZ x1220 = x808 * x442;
-Tbase TZ x1221 = x1219 + x1220;
-Tbase TZ x1222 = x1216 + x1221;
-Tbase TZ x1223 = 0x2;
-Tbase TZ x1224 = x439 * x1223;
-Tbase TZ x1225 = x799 * x1224;
-Tbase TZ x1226 = x800 * x436;
-Tbase TZ x1227 = 0x2;
-Tbase TZ x1228 = x433 * x1227;
-Tbase TZ x1229 = x801 * x1228;
-Tbase TZ x1230 = x802 * x430;
-Tbase TZ x1231 = 0x2;
-Tbase TZ x1232 = x427 * x1231;
-Tbase TZ x1233 = x803 * x1232;
-Tbase TZ x1234 = x804 * x424;
-Tbase TZ x1235 = 0x2;
-Tbase TZ x1236 = x421 * x1235;
-Tbase TZ x1237 = x805 * x1236;
-Tbase TZ x1238 = x1234 + x1237;
-Tbase TZ x1239 = x1233 + x1238;
-Tbase TZ x1240 = x1230 + x1239;
-Tbase TZ x1241 = x1229 + x1240;
-Tbase TZ x1242 = x1226 + x1241;
-Tbase TZ x1243 = x1225 + x1242;
-Tbase TZ x1244 = 0x13;
-Tbase TZ x1245 = x1244 * x1243;
-Tbase TZ x1246 = x1222 + x1245;
-Tbase TZ x1247 = x1215 + x1246;
-Tbase TZ x1248 = 0x1a;
-Tbase TZ x1249 = x1247 >> x1248;
-Tbase TZ x1250 = x805 * x448;
-Tbase TZ x1251 = x806 * x445;
-Tbase TZ x1252 = x807 * x442;
-Tbase TZ x1253 = x808 * x439;
-Tbase TZ x1254 = x1252 + x1253;
-Tbase TZ x1255 = x1251 + x1254;
-Tbase TZ x1256 = x1250 + x1255;
-Tbase TZ x1257 = x799 * x436;
-Tbase TZ x1258 = x800 * x433;
-Tbase TZ x1259 = x801 * x430;
-Tbase TZ x1260 = x802 * x427;
-Tbase TZ x1261 = x803 * x424;
-Tbase TZ x1262 = x804 * x421;
-Tbase TZ x1263 = x1261 + x1262;
-Tbase TZ x1264 = x1260 + x1263;
-Tbase TZ x1265 = x1259 + x1264;
-Tbase TZ x1266 = x1258 + x1265;
-Tbase TZ x1267 = x1257 + x1266;
-Tbase TZ x1268 = 0x13;
-Tbase TZ x1269 = x1268 * x1267;
-Tbase TZ x1270 = x1256 + x1269;
-Tbase TZ x1271 = x1249 + x1270;
-Tbase TZ x1272 = 0x19;
-Tbase TZ x1273 = x1271 >> x1272;
-Tbase TZ x1274 = x804 * x448;
-Tbase TZ x1275 = 0x2;
-Tbase TZ x1276 = x445 * x1275;
-Tbase TZ x1277 = x805 * x1276;
-Tbase TZ x1278 = x806 * x442;
-Tbase TZ x1279 = 0x2;
-Tbase TZ x1280 = x439 * x1279;
-Tbase TZ x1281 = x807 * x1280;
-Tbase TZ x1282 = x808 * x436;
-Tbase TZ x1283 = x1281 + x1282;
-Tbase TZ x1284 = x1278 + x1283;
-Tbase TZ x1285 = x1277 + x1284;
-Tbase TZ x1286 = x1274 + x1285;
-Tbase TZ x1287 = 0x2;
-Tbase TZ x1288 = x433 * x1287;
-Tbase TZ x1289 = x799 * x1288;
-Tbase TZ x1290 = x800 * x430;
-Tbase TZ x1291 = 0x2;
-Tbase TZ x1292 = x427 * x1291;
-Tbase TZ x1293 = x801 * x1292;
-Tbase TZ x1294 = x802 * x424;
-Tbase TZ x1295 = 0x2;
-Tbase TZ x1296 = x421 * x1295;
-Tbase TZ x1297 = x803 * x1296;
-Tbase TZ x1298 = x1294 + x1297;
-Tbase TZ x1299 = x1293 + x1298;
-Tbase TZ x1300 = x1290 + x1299;
-Tbase TZ x1301 = x1289 + x1300;
-Tbase TZ x1302 = 0x13;
-Tbase TZ x1303 = x1302 * x1301;
-Tbase TZ x1304 = x1286 + x1303;
-Tbase TZ x1305 = x1273 + x1304;
-Tbase TZ x1306 = 0x1a;
-Tbase TZ x1307 = x1305 >> x1306;
-Tbase TZ x1308 = x803 * x448;
-Tbase TZ x1309 = x804 * x445;
-Tbase TZ x1310 = x805 * x442;
-Tbase TZ x1311 = x806 * x439;
-Tbase TZ x1312 = x807 * x436;
-Tbase TZ x1313 = x808 * x433;
-Tbase TZ x1314 = x1312 + x1313;
-Tbase TZ x1315 = x1311 + x1314;
-Tbase TZ x1316 = x1310 + x1315;
-Tbase TZ x1317 = x1309 + x1316;
-Tbase TZ x1318 = x1308 + x1317;
-Tbase TZ x1319 = x799 * x430;
-Tbase TZ x1320 = x800 * x427;
-Tbase TZ x1321 = x801 * x424;
-Tbase TZ x1322 = x802 * x421;
-Tbase TZ x1323 = x1321 + x1322;
-Tbase TZ x1324 = x1320 + x1323;
-Tbase TZ x1325 = x1319 + x1324;
-Tbase TZ x1326 = 0x13;
-Tbase TZ x1327 = x1326 * x1325;
-Tbase TZ x1328 = x1318 + x1327;
-Tbase TZ x1329 = x1307 + x1328;
-Tbase TZ x1330 = 0x19;
-Tbase TZ x1331 = x1329 >> x1330;
-Tbase TZ x1332 = x802 * x448;
-Tbase TZ x1333 = 0x2;
-Tbase TZ x1334 = x445 * x1333;
-Tbase TZ x1335 = x803 * x1334;
-Tbase TZ x1336 = x804 * x442;
-Tbase TZ x1337 = 0x2;
-Tbase TZ x1338 = x439 * x1337;
-Tbase TZ x1339 = x805 * x1338;
-Tbase TZ x1340 = x806 * x436;
-Tbase TZ x1341 = 0x2;
-Tbase TZ x1342 = x433 * x1341;
-Tbase TZ x1343 = x807 * x1342;
-Tbase TZ x1344 = x808 * x430;
-Tbase TZ x1345 = x1343 + x1344;
-Tbase TZ x1346 = x1340 + x1345;
-Tbase TZ x1347 = x1339 + x1346;
-Tbase TZ x1348 = x1336 + x1347;
-Tbase TZ x1349 = x1335 + x1348;
-Tbase TZ x1350 = x1332 + x1349;
-Tbase TZ x1351 = 0x2;
-Tbase TZ x1352 = x427 * x1351;
-Tbase TZ x1353 = x799 * x1352;
-Tbase TZ x1354 = x800 * x424;
-Tbase TZ x1355 = 0x2;
-Tbase TZ x1356 = x421 * x1355;
-Tbase TZ x1357 = x801 * x1356;
-Tbase TZ x1358 = x1354 + x1357;
-Tbase TZ x1359 = x1353 + x1358;
-Tbase TZ x1360 = 0x13;
-Tbase TZ x1361 = x1360 * x1359;
-Tbase TZ x1362 = x1350 + x1361;
-Tbase TZ x1363 = x1331 + x1362;
-Tbase TZ x1364 = 0x1a;
-Tbase TZ x1365 = x1363 >> x1364;
-Tbase TZ x1366 = x801 * x448;
-Tbase TZ x1367 = x802 * x445;
-Tbase TZ x1368 = x803 * x442;
-Tbase TZ x1369 = x804 * x439;
-Tbase TZ x1370 = x805 * x436;
-Tbase TZ x1371 = x806 * x433;
-Tbase TZ x1372 = x807 * x430;
-Tbase TZ x1373 = x808 * x427;
-Tbase TZ x1374 = x1372 + x1373;
-Tbase TZ x1375 = x1371 + x1374;
-Tbase TZ x1376 = x1370 + x1375;
-Tbase TZ x1377 = x1369 + x1376;
-Tbase TZ x1378 = x1368 + x1377;
-Tbase TZ x1379 = x1367 + x1378;
-Tbase TZ x1380 = x1366 + x1379;
-Tbase TZ x1381 = x799 * x424;
-Tbase TZ x1382 = x800 * x421;
-Tbase TZ x1383 = x1381 + x1382;
-Tbase TZ x1384 = 0x13;
-Tbase TZ x1385 = x1384 * x1383;
-Tbase TZ x1386 = x1380 + x1385;
-Tbase TZ x1387 = x1365 + x1386;
-Tbase TZ x1388 = 0x19;
-Tbase TZ x1389 = x1387 >> x1388;
-Tbase TZ x1390 = x800 * x448;
-Tbase TZ x1391 = 0x2;
-Tbase TZ x1392 = x445 * x1391;
-Tbase TZ x1393 = x801 * x1392;
-Tbase TZ x1394 = x802 * x442;
-Tbase TZ x1395 = 0x2;
-Tbase TZ x1396 = x439 * x1395;
-Tbase TZ x1397 = x803 * x1396;
-Tbase TZ x1398 = x804 * x436;
-Tbase TZ x1399 = 0x2;
-Tbase TZ x1400 = x433 * x1399;
-Tbase TZ x1401 = x805 * x1400;
-Tbase TZ x1402 = x806 * x430;
-Tbase TZ x1403 = 0x2;
-Tbase TZ x1404 = x427 * x1403;
-Tbase TZ x1405 = x807 * x1404;
-Tbase TZ x1406 = x808 * x424;
-Tbase TZ x1407 = x1405 + x1406;
-Tbase TZ x1408 = x1402 + x1407;
-Tbase TZ x1409 = x1401 + x1408;
-Tbase TZ x1410 = x1398 + x1409;
-Tbase TZ x1411 = x1397 + x1410;
-Tbase TZ x1412 = x1394 + x1411;
-Tbase TZ x1413 = x1393 + x1412;
-Tbase TZ x1414 = x1390 + x1413;
-Tbase TZ x1415 = 0x2;
-Tbase TZ x1416 = x421 * x1415;
-Tbase TZ x1417 = x799 * x1416;
-Tbase TZ x1418 = 0x13;
-Tbase TZ x1419 = x1418 * x1417;
-Tbase TZ x1420 = x1414 + x1419;
-Tbase TZ x1421 = x1389 + x1420;
-Tbase TZ x1422 = 0x1a;
-Tbase TZ x1423 = x1421 >> x1422;
-Tbase TZ x1424 = x799 * x448;
-Tbase TZ x1425 = x800 * x445;
-Tbase TZ x1426 = x801 * x442;
-Tbase TZ x1427 = x802 * x439;
-Tbase TZ x1428 = x803 * x436;
-Tbase TZ x1429 = x804 * x433;
-Tbase TZ x1430 = x805 * x430;
-Tbase TZ x1431 = x806 * x427;
-Tbase TZ x1432 = x807 * x424;
-Tbase TZ x1433 = x808 * x421;
-Tbase TZ x1434 = x1432 + x1433;
-Tbase TZ x1435 = x1431 + x1434;
-Tbase TZ x1436 = x1430 + x1435;
-Tbase TZ x1437 = x1429 + x1436;
-Tbase TZ x1438 = x1428 + x1437;
-Tbase TZ x1439 = x1427 + x1438;
-Tbase TZ x1440 = x1426 + x1439;
-Tbase TZ x1441 = x1425 + x1440;
-Tbase TZ x1442 = x1424 + x1441;
-Tbase TZ x1443 = x1423 + x1442;
-Tbase TZ x1444 = 0x19;
-Tbase TZ x1445 = x1443 >> x1444;
-Tbase TZ x1446 = 0x13;
-Tbase TZ x1447 = x1446 * x1445;
-Tbase TZ x1448 = 0x3ffffff;
-Tbase TZ x1449 = x1189 & x1448;
-Tbase TZ x1450 = x1447 + x1449;
-Tbase TZ x1451 = 0x1a;
-Tbase TZ x1452 = x1450 >> x1451;
-Tbase TZ x1453 = 0x1ffffff;
-Tbase TZ x1454 = x1213 & x1453;
-Tbase TZ x1455 = x1452 + x1454;
-Tbase TZ x1456 = 0x1ffffff;
-Tbase TZ x1457 = x1443 & x1456;
-Tbase TZ x1458 = 0x3ffffff;
-Tbase TZ x1459 = x1421 & x1458;
-Tbase TZ x1460 = 0x1ffffff;
-Tbase TZ x1461 = x1387 & x1460;
-Tbase TZ x1462 = 0x3ffffff;
-Tbase TZ x1463 = x1363 & x1462;
-Tbase TZ x1464 = 0x1ffffff;
-Tbase TZ x1465 = x1329 & x1464;
-Tbase TZ x1466 = 0x3ffffff;
-Tbase TZ x1467 = x1305 & x1466;
-Tbase TZ x1468 = 0x1ffffff;
-Tbase TZ x1469 = x1271 & x1468;
-Tbase TZ x1470 = 0x19;
-Tbase TZ x1471 = x1455 >> x1470;
-Tbase TZ x1472 = 0x3ffffff;
-Tbase TZ x1473 = x1247 & x1472;
-Tbase TZ x1474 = x1471 + x1473;
-Tbase TZ x1475 = 0x1ffffff;
-Tbase TZ x1476 = x1455 & x1475;
-Tbase TZ x1477 = 0x3ffffff;
-Tbase TZ x1478 = x1450 & x1477;
-Tbase TZ x1479 = x1137 + x1457;
-Tbase TZ x1480 = x1139 + x1459;
-Tbase TZ x1481 = x1141 + x1461;
-Tbase TZ x1482 = x1143 + x1463;
-Tbase TZ x1483 = x1145 + x1465;
-Tbase TZ x1484 = x1147 + x1467;
-Tbase TZ x1485 = x1149 + x1469;
-Tbase TZ x1486 = x1154 + x1474;
-Tbase TZ x1487 = x1156 + x1476;
-Tbase TZ x1488 = x1158 + x1478;
-Tbase TZ x1489 = x1137 + x1457;
-Tbase TZ x1490 = x1139 + x1459;
-Tbase TZ x1491 = x1141 + x1461;
-Tbase TZ x1492 = x1143 + x1463;
-Tbase TZ x1493 = x1145 + x1465;
-Tbase TZ x1494 = x1147 + x1467;
-Tbase TZ x1495 = x1149 + x1469;
-Tbase TZ x1496 = x1154 + x1474;
-Tbase TZ x1497 = x1156 + x1476;
-Tbase TZ x1498 = x1158 + x1478;
-Tbase TZ x1499 = x1488 * x1498;
-Tbase TZ x1500 = 0x2;
-Tbase TZ x1501 = x1497 * x1500;
-Tbase TZ x1502 = x1479 * x1501;
-Tbase TZ x1503 = x1480 * x1496;
-Tbase TZ x1504 = 0x2;
-Tbase TZ x1505 = x1495 * x1504;
-Tbase TZ x1506 = x1481 * x1505;
-Tbase TZ x1507 = x1482 * x1494;
-Tbase TZ x1508 = 0x2;
-Tbase TZ x1509 = x1493 * x1508;
-Tbase TZ x1510 = x1483 * x1509;
-Tbase TZ x1511 = x1484 * x1492;
-Tbase TZ x1512 = 0x2;
-Tbase TZ x1513 = x1491 * x1512;
-Tbase TZ x1514 = x1485 * x1513;
-Tbase TZ x1515 = x1486 * x1490;
-Tbase TZ x1516 = 0x2;
-Tbase TZ x1517 = x1489 * x1516;
-Tbase TZ x1518 = x1487 * x1517;
-Tbase TZ x1519 = x1515 + x1518;
-Tbase TZ x1520 = x1514 + x1519;
-Tbase TZ x1521 = x1511 + x1520;
-Tbase TZ x1522 = x1510 + x1521;
-Tbase TZ x1523 = x1507 + x1522;
-Tbase TZ x1524 = x1506 + x1523;
-Tbase TZ x1525 = x1503 + x1524;
-Tbase TZ x1526 = x1502 + x1525;
-Tbase TZ x1527 = 0x13;
-Tbase TZ x1528 = x1527 * x1526;
-Tbase TZ x1529 = x1499 + x1528;
-Tbase TZ x1530 = 0x1a;
-Tbase TZ x1531 = x1529 >> x1530;
-Tbase TZ x1532 = x1487 * x1498;
-Tbase TZ x1533 = x1488 * x1497;
-Tbase TZ x1534 = x1532 + x1533;
-Tbase TZ x1535 = x1479 * x1496;
-Tbase TZ x1536 = x1480 * x1495;
-Tbase TZ x1537 = x1481 * x1494;
-Tbase TZ x1538 = x1482 * x1493;
-Tbase TZ x1539 = x1483 * x1492;
-Tbase TZ x1540 = x1484 * x1491;
-Tbase TZ x1541 = x1485 * x1490;
-Tbase TZ x1542 = x1486 * x1489;
-Tbase TZ x1543 = x1541 + x1542;
-Tbase TZ x1544 = x1540 + x1543;
-Tbase TZ x1545 = x1539 + x1544;
-Tbase TZ x1546 = x1538 + x1545;
-Tbase TZ x1547 = x1537 + x1546;
-Tbase TZ x1548 = x1536 + x1547;
-Tbase TZ x1549 = x1535 + x1548;
-Tbase TZ x1550 = 0x13;
-Tbase TZ x1551 = x1550 * x1549;
-Tbase TZ x1552 = x1534 + x1551;
-Tbase TZ x1553 = x1531 + x1552;
-Tbase TZ x1554 = 0x19;
-Tbase TZ x1555 = x1553 >> x1554;
-Tbase TZ x1556 = x1486 * x1498;
-Tbase TZ x1557 = 0x2;
-Tbase TZ x1558 = x1497 * x1557;
-Tbase TZ x1559 = x1487 * x1558;
-Tbase TZ x1560 = x1488 * x1496;
-Tbase TZ x1561 = x1559 + x1560;
-Tbase TZ x1562 = x1556 + x1561;
-Tbase TZ x1563 = 0x2;
-Tbase TZ x1564 = x1495 * x1563;
-Tbase TZ x1565 = x1479 * x1564;
-Tbase TZ x1566 = x1480 * x1494;
-Tbase TZ x1567 = 0x2;
-Tbase TZ x1568 = x1493 * x1567;
-Tbase TZ x1569 = x1481 * x1568;
-Tbase TZ x1570 = x1482 * x1492;
-Tbase TZ x1571 = 0x2;
-Tbase TZ x1572 = x1491 * x1571;
-Tbase TZ x1573 = x1483 * x1572;
-Tbase TZ x1574 = x1484 * x1490;
-Tbase TZ x1575 = 0x2;
-Tbase TZ x1576 = x1489 * x1575;
-Tbase TZ x1577 = x1485 * x1576;
-Tbase TZ x1578 = x1574 + x1577;
-Tbase TZ x1579 = x1573 + x1578;
-Tbase TZ x1580 = x1570 + x1579;
-Tbase TZ x1581 = x1569 + x1580;
-Tbase TZ x1582 = x1566 + x1581;
-Tbase TZ x1583 = x1565 + x1582;
-Tbase TZ x1584 = 0x13;
-Tbase TZ x1585 = x1584 * x1583;
-Tbase TZ x1586 = x1562 + x1585;
-Tbase TZ x1587 = x1555 + x1586;
-Tbase TZ x1588 = 0x1a;
-Tbase TZ x1589 = x1587 >> x1588;
-Tbase TZ x1590 = x1485 * x1498;
-Tbase TZ x1591 = x1486 * x1497;
-Tbase TZ x1592 = x1487 * x1496;
-Tbase TZ x1593 = x1488 * x1495;
-Tbase TZ x1594 = x1592 + x1593;
-Tbase TZ x1595 = x1591 + x1594;
-Tbase TZ x1596 = x1590 + x1595;
-Tbase TZ x1597 = x1479 * x1494;
-Tbase TZ x1598 = x1480 * x1493;
-Tbase TZ x1599 = x1481 * x1492;
-Tbase TZ x1600 = x1482 * x1491;
-Tbase TZ x1601 = x1483 * x1490;
-Tbase TZ x1602 = x1484 * x1489;
-Tbase TZ x1603 = x1601 + x1602;
-Tbase TZ x1604 = x1600 + x1603;
-Tbase TZ x1605 = x1599 + x1604;
-Tbase TZ x1606 = x1598 + x1605;
-Tbase TZ x1607 = x1597 + x1606;
-Tbase TZ x1608 = 0x13;
-Tbase TZ x1609 = x1608 * x1607;
-Tbase TZ x1610 = x1596 + x1609;
-Tbase TZ x1611 = x1589 + x1610;
-Tbase TZ x1612 = 0x19;
-Tbase TZ x1613 = x1611 >> x1612;
-Tbase TZ x1614 = x1484 * x1498;
-Tbase TZ x1615 = 0x2;
-Tbase TZ x1616 = x1497 * x1615;
-Tbase TZ x1617 = x1485 * x1616;
-Tbase TZ x1618 = x1486 * x1496;
-Tbase TZ x1619 = 0x2;
-Tbase TZ x1620 = x1495 * x1619;
-Tbase TZ x1621 = x1487 * x1620;
-Tbase TZ x1622 = x1488 * x1494;
-Tbase TZ x1623 = x1621 + x1622;
-Tbase TZ x1624 = x1618 + x1623;
-Tbase TZ x1625 = x1617 + x1624;
-Tbase TZ x1626 = x1614 + x1625;
-Tbase TZ x1627 = 0x2;
-Tbase TZ x1628 = x1493 * x1627;
-Tbase TZ x1629 = x1479 * x1628;
-Tbase TZ x1630 = x1480 * x1492;
-Tbase TZ x1631 = 0x2;
-Tbase TZ x1632 = x1491 * x1631;
-Tbase TZ x1633 = x1481 * x1632;
-Tbase TZ x1634 = x1482 * x1490;
-Tbase TZ x1635 = 0x2;
-Tbase TZ x1636 = x1489 * x1635;
-Tbase TZ x1637 = x1483 * x1636;
-Tbase TZ x1638 = x1634 + x1637;
-Tbase TZ x1639 = x1633 + x1638;
-Tbase TZ x1640 = x1630 + x1639;
-Tbase TZ x1641 = x1629 + x1640;
-Tbase TZ x1642 = 0x13;
-Tbase TZ x1643 = x1642 * x1641;
-Tbase TZ x1644 = x1626 + x1643;
-Tbase TZ x1645 = x1613 + x1644;
-Tbase TZ x1646 = 0x1a;
-Tbase TZ x1647 = x1645 >> x1646;
-Tbase TZ x1648 = x1483 * x1498;
-Tbase TZ x1649 = x1484 * x1497;
-Tbase TZ x1650 = x1485 * x1496;
-Tbase TZ x1651 = x1486 * x1495;
-Tbase TZ x1652 = x1487 * x1494;
-Tbase TZ x1653 = x1488 * x1493;
-Tbase TZ x1654 = x1652 + x1653;
-Tbase TZ x1655 = x1651 + x1654;
-Tbase TZ x1656 = x1650 + x1655;
-Tbase TZ x1657 = x1649 + x1656;
-Tbase TZ x1658 = x1648 + x1657;
-Tbase TZ x1659 = x1479 * x1492;
-Tbase TZ x1660 = x1480 * x1491;
-Tbase TZ x1661 = x1481 * x1490;
-Tbase TZ x1662 = x1482 * x1489;
-Tbase TZ x1663 = x1661 + x1662;
-Tbase TZ x1664 = x1660 + x1663;
-Tbase TZ x1665 = x1659 + x1664;
-Tbase TZ x1666 = 0x13;
-Tbase TZ x1667 = x1666 * x1665;
-Tbase TZ x1668 = x1658 + x1667;
-Tbase TZ x1669 = x1647 + x1668;
-Tbase TZ x1670 = 0x19;
-Tbase TZ x1671 = x1669 >> x1670;
-Tbase TZ x1672 = x1482 * x1498;
-Tbase TZ x1673 = 0x2;
-Tbase TZ x1674 = x1497 * x1673;
-Tbase TZ x1675 = x1483 * x1674;
-Tbase TZ x1676 = x1484 * x1496;
-Tbase TZ x1677 = 0x2;
-Tbase TZ x1678 = x1495 * x1677;
-Tbase TZ x1679 = x1485 * x1678;
-Tbase TZ x1680 = x1486 * x1494;
-Tbase TZ x1681 = 0x2;
-Tbase TZ x1682 = x1493 * x1681;
-Tbase TZ x1683 = x1487 * x1682;
-Tbase TZ x1684 = x1488 * x1492;
-Tbase TZ x1685 = x1683 + x1684;
-Tbase TZ x1686 = x1680 + x1685;
-Tbase TZ x1687 = x1679 + x1686;
-Tbase TZ x1688 = x1676 + x1687;
-Tbase TZ x1689 = x1675 + x1688;
-Tbase TZ x1690 = x1672 + x1689;
-Tbase TZ x1691 = 0x2;
-Tbase TZ x1692 = x1491 * x1691;
-Tbase TZ x1693 = x1479 * x1692;
-Tbase TZ x1694 = x1480 * x1490;
-Tbase TZ x1695 = 0x2;
-Tbase TZ x1696 = x1489 * x1695;
-Tbase TZ x1697 = x1481 * x1696;
-Tbase TZ x1698 = x1694 + x1697;
-Tbase TZ x1699 = x1693 + x1698;
-Tbase TZ x1700 = 0x13;
-Tbase TZ x1701 = x1700 * x1699;
-Tbase TZ x1702 = x1690 + x1701;
-Tbase TZ x1703 = x1671 + x1702;
-Tbase TZ x1704 = 0x1a;
-Tbase TZ x1705 = x1703 >> x1704;
-Tbase TZ x1706 = x1481 * x1498;
-Tbase TZ x1707 = x1482 * x1497;
-Tbase TZ x1708 = x1483 * x1496;
-Tbase TZ x1709 = x1484 * x1495;
-Tbase TZ x1710 = x1485 * x1494;
-Tbase TZ x1711 = x1486 * x1493;
-Tbase TZ x1712 = x1487 * x1492;
-Tbase TZ x1713 = x1488 * x1491;
-Tbase TZ x1714 = x1712 + x1713;
-Tbase TZ x1715 = x1711 + x1714;
-Tbase TZ x1716 = x1710 + x1715;
-Tbase TZ x1717 = x1709 + x1716;
-Tbase TZ x1718 = x1708 + x1717;
-Tbase TZ x1719 = x1707 + x1718;
-Tbase TZ x1720 = x1706 + x1719;
-Tbase TZ x1721 = x1479 * x1490;
-Tbase TZ x1722 = x1480 * x1489;
-Tbase TZ x1723 = x1721 + x1722;
-Tbase TZ x1724 = 0x13;
-Tbase TZ x1725 = x1724 * x1723;
-Tbase TZ x1726 = x1720 + x1725;
-Tbase TZ x1727 = x1705 + x1726;
-Tbase TZ x1728 = 0x19;
-Tbase TZ x1729 = x1727 >> x1728;
-Tbase TZ x1730 = x1480 * x1498;
-Tbase TZ x1731 = 0x2;
-Tbase TZ x1732 = x1497 * x1731;
-Tbase TZ x1733 = x1481 * x1732;
-Tbase TZ x1734 = x1482 * x1496;
-Tbase TZ x1735 = 0x2;
-Tbase TZ x1736 = x1495 * x1735;
-Tbase TZ x1737 = x1483 * x1736;
-Tbase TZ x1738 = x1484 * x1494;
-Tbase TZ x1739 = 0x2;
-Tbase TZ x1740 = x1493 * x1739;
-Tbase TZ x1741 = x1485 * x1740;
-Tbase TZ x1742 = x1486 * x1492;
-Tbase TZ x1743 = 0x2;
-Tbase TZ x1744 = x1491 * x1743;
-Tbase TZ x1745 = x1487 * x1744;
-Tbase TZ x1746 = x1488 * x1490;
-Tbase TZ x1747 = x1745 + x1746;
-Tbase TZ x1748 = x1742 + x1747;
-Tbase TZ x1749 = x1741 + x1748;
-Tbase TZ x1750 = x1738 + x1749;
-Tbase TZ x1751 = x1737 + x1750;
-Tbase TZ x1752 = x1734 + x1751;
-Tbase TZ x1753 = x1733 + x1752;
-Tbase TZ x1754 = x1730 + x1753;
-Tbase TZ x1755 = 0x2;
-Tbase TZ x1756 = x1489 * x1755;
-Tbase TZ x1757 = x1479 * x1756;
-Tbase TZ x1758 = 0x13;
-Tbase TZ x1759 = x1758 * x1757;
-Tbase TZ x1760 = x1754 + x1759;
-Tbase TZ x1761 = x1729 + x1760;
-Tbase TZ x1762 = 0x1a;
-Tbase TZ x1763 = x1761 >> x1762;
-Tbase TZ x1764 = x1479 * x1498;
-Tbase TZ x1765 = x1480 * x1497;
-Tbase TZ x1766 = x1481 * x1496;
-Tbase TZ x1767 = x1482 * x1495;
-Tbase TZ x1768 = x1483 * x1494;
-Tbase TZ x1769 = x1484 * x1493;
-Tbase TZ x1770 = x1485 * x1492;
-Tbase TZ x1771 = x1486 * x1491;
-Tbase TZ x1772 = x1487 * x1490;
-Tbase TZ x1773 = x1488 * x1489;
-Tbase TZ x1774 = x1772 + x1773;
-Tbase TZ x1775 = x1771 + x1774;
-Tbase TZ x1776 = x1770 + x1775;
-Tbase TZ x1777 = x1769 + x1776;
-Tbase TZ x1778 = x1768 + x1777;
-Tbase TZ x1779 = x1767 + x1778;
-Tbase TZ x1780 = x1766 + x1779;
-Tbase TZ x1781 = x1765 + x1780;
-Tbase TZ x1782 = x1764 + x1781;
-Tbase TZ x1783 = x1763 + x1782;
-Tbase TZ x1784 = 0x19;
-Tbase TZ x1785 = x1783 >> x1784;
-Tbase TZ x1786 = 0x13;
-Tbase TZ x1787 = x1786 * x1785;
-Tbase TZ x1788 = 0x3ffffff;
-Tbase TZ x1789 = x1529 & x1788;
-Tbase TZ x1790 = x1787 + x1789;
-Tbase TZ x1791 = 0x1a;
-Tbase TZ x1792 = x1790 >> x1791;
-Tbase TZ x1793 = 0x1ffffff;
-Tbase TZ x1794 = x1553 & x1793;
-Tbase TZ x1795 = x1792 + x1794;
-Tbase TZ x1796 = 0x1ffffff;
-Tbase TZ x1797 = x1783 & x1796;
-Tbase TZ x1798 = 0x3ffffff;
-Tbase TZ x1799 = x1761 & x1798;
-Tbase TZ x1800 = 0x1ffffff;
-Tbase TZ x1801 = x1727 & x1800;
-Tbase TZ x1802 = 0x3ffffff;
-Tbase TZ x1803 = x1703 & x1802;
-Tbase TZ x1804 = 0x1ffffff;
-Tbase TZ x1805 = x1669 & x1804;
-Tbase TZ x1806 = 0x3ffffff;
-Tbase TZ x1807 = x1645 & x1806;
-Tbase TZ x1808 = 0x1ffffff;
-Tbase TZ x1809 = x1611 & x1808;
-Tbase TZ x1810 = 0x19;
-Tbase TZ x1811 = x1795 >> x1810;
-Tbase TZ x1812 = 0x3ffffff;
-Tbase TZ x1813 = x1587 & x1812;
-Tbase TZ x1814 = x1811 + x1813;
-Tbase TZ x1815 = 0x1ffffff;
-Tbase TZ x1816 = x1795 & x1815;
-Tbase TZ x1817 = 0x3ffffff;
-Tbase TZ x1818 = x1790 & x1817;
-Tbase TZ x1819 = 0x3fffffe;
-Tbase TZ x1820 = x1819 + x1137;
-Tbase TZ x1821 = x1820 - x1457;
-Tbase TZ x1822 = 0x7fffffe;
-Tbase TZ x1823 = x1822 + x1139;
-Tbase TZ x1824 = x1823 - x1459;
-Tbase TZ x1825 = 0x3fffffe;
-Tbase TZ x1826 = x1825 + x1141;
-Tbase TZ x1827 = x1826 - x1461;
-Tbase TZ x1828 = 0x7fffffe;
-Tbase TZ x1829 = x1828 + x1143;
-Tbase TZ x1830 = x1829 - x1463;
-Tbase TZ x1831 = 0x3fffffe;
-Tbase TZ x1832 = x1831 + x1145;
-Tbase TZ x1833 = x1832 - x1465;
-Tbase TZ x1834 = 0x7fffffe;
-Tbase TZ x1835 = x1834 + x1147;
-Tbase TZ x1836 = x1835 - x1467;
-Tbase TZ x1837 = 0x3fffffe;
-Tbase TZ x1838 = x1837 + x1149;
-Tbase TZ x1839 = x1838 - x1469;
-Tbase TZ x1840 = 0x7fffffe;
-Tbase TZ x1841 = x1840 + x1154;
-Tbase TZ x1842 = x1841 - x1474;
-Tbase TZ x1843 = 0x3fffffe;
-Tbase TZ x1844 = x1843 + x1156;
-Tbase TZ x1845 = x1844 - x1476;
-Tbase TZ x1846 = 0x7ffffda;
-Tbase TZ x1847 = x1846 + x1158;
-Tbase TZ x1848 = x1847 - x1478;
-Tbase TZ x1849 = 0x3fffffe;
-Tbase TZ x1850 = x1849 + x1137;
-Tbase TZ x1851 = x1850 - x1457;
-Tbase TZ x1852 = 0x7fffffe;
-Tbase TZ x1853 = x1852 + x1139;
-Tbase TZ x1854 = x1853 - x1459;
-Tbase TZ x1855 = 0x3fffffe;
-Tbase TZ x1856 = x1855 + x1141;
-Tbase TZ x1857 = x1856 - x1461;
-Tbase TZ x1858 = 0x7fffffe;
-Tbase TZ x1859 = x1858 + x1143;
-Tbase TZ x1860 = x1859 - x1463;
-Tbase TZ x1861 = 0x3fffffe;
-Tbase TZ x1862 = x1861 + x1145;
-Tbase TZ x1863 = x1862 - x1465;
-Tbase TZ x1864 = 0x7fffffe;
-Tbase TZ x1865 = x1864 + x1147;
-Tbase TZ x1866 = x1865 - x1467;
-Tbase TZ x1867 = 0x3fffffe;
-Tbase TZ x1868 = x1867 + x1149;
-Tbase TZ x1869 = x1868 - x1469;
-Tbase TZ x1870 = 0x7fffffe;
-Tbase TZ x1871 = x1870 + x1154;
-Tbase TZ x1872 = x1871 - x1474;
-Tbase TZ x1873 = 0x3fffffe;
-Tbase TZ x1874 = x1873 + x1156;
-Tbase TZ x1875 = x1874 - x1476;
-Tbase TZ x1876 = 0x7ffffda;
-Tbase TZ x1877 = x1876 + x1158;
-Tbase TZ x1878 = x1877 - x1478;
-Tbase TZ x1879 = x1848 * x1878;
-Tbase TZ x1880 = 0x2;
-Tbase TZ x1881 = x1875 * x1880;
-Tbase TZ x1882 = x1821 * x1881;
-Tbase TZ x1883 = x1824 * x1872;
-Tbase TZ x1884 = 0x2;
-Tbase TZ x1885 = x1869 * x1884;
-Tbase TZ x1886 = x1827 * x1885;
-Tbase TZ x1887 = x1830 * x1866;
-Tbase TZ x1888 = 0x2;
-Tbase TZ x1889 = x1863 * x1888;
-Tbase TZ x1890 = x1833 * x1889;
-Tbase TZ x1891 = x1836 * x1860;
-Tbase TZ x1892 = 0x2;
-Tbase TZ x1893 = x1857 * x1892;
-Tbase TZ x1894 = x1839 * x1893;
-Tbase TZ x1895 = x1842 * x1854;
-Tbase TZ x1896 = 0x2;
-Tbase TZ x1897 = x1851 * x1896;
-Tbase TZ x1898 = x1845 * x1897;
-Tbase TZ x1899 = x1895 + x1898;
-Tbase TZ x1900 = x1894 + x1899;
-Tbase TZ x1901 = x1891 + x1900;
-Tbase TZ x1902 = x1890 + x1901;
-Tbase TZ x1903 = x1887 + x1902;
-Tbase TZ x1904 = x1886 + x1903;
-Tbase TZ x1905 = x1883 + x1904;
-Tbase TZ x1906 = x1882 + x1905;
-Tbase TZ x1907 = 0x13;
-Tbase TZ x1908 = x1907 * x1906;
-Tbase TZ x1909 = x1879 + x1908;
-Tbase TZ x1910 = 0x1a;
-Tbase TZ x1911 = x1909 >> x1910;
-Tbase TZ x1912 = x1845 * x1878;
-Tbase TZ x1913 = x1848 * x1875;
-Tbase TZ x1914 = x1912 + x1913;
-Tbase TZ x1915 = x1821 * x1872;
-Tbase TZ x1916 = x1824 * x1869;
-Tbase TZ x1917 = x1827 * x1866;
-Tbase TZ x1918 = x1830 * x1863;
-Tbase TZ x1919 = x1833 * x1860;
-Tbase TZ x1920 = x1836 * x1857;
-Tbase TZ x1921 = x1839 * x1854;
-Tbase TZ x1922 = x1842 * x1851;
-Tbase TZ x1923 = x1921 + x1922;
-Tbase TZ x1924 = x1920 + x1923;
-Tbase TZ x1925 = x1919 + x1924;
-Tbase TZ x1926 = x1918 + x1925;
-Tbase TZ x1927 = x1917 + x1926;
-Tbase TZ x1928 = x1916 + x1927;
-Tbase TZ x1929 = x1915 + x1928;
-Tbase TZ x1930 = 0x13;
-Tbase TZ x1931 = x1930 * x1929;
-Tbase TZ x1932 = x1914 + x1931;
-Tbase TZ x1933 = x1911 + x1932;
-Tbase TZ x1934 = 0x19;
-Tbase TZ x1935 = x1933 >> x1934;
-Tbase TZ x1936 = x1842 * x1878;
-Tbase TZ x1937 = 0x2;
-Tbase TZ x1938 = x1875 * x1937;
-Tbase TZ x1939 = x1845 * x1938;
-Tbase TZ x1940 = x1848 * x1872;
-Tbase TZ x1941 = x1939 + x1940;
-Tbase TZ x1942 = x1936 + x1941;
-Tbase TZ x1943 = 0x2;
-Tbase TZ x1944 = x1869 * x1943;
-Tbase TZ x1945 = x1821 * x1944;
-Tbase TZ x1946 = x1824 * x1866;
-Tbase TZ x1947 = 0x2;
-Tbase TZ x1948 = x1863 * x1947;
-Tbase TZ x1949 = x1827 * x1948;
-Tbase TZ x1950 = x1830 * x1860;
-Tbase TZ x1951 = 0x2;
-Tbase TZ x1952 = x1857 * x1951;
-Tbase TZ x1953 = x1833 * x1952;
-Tbase TZ x1954 = x1836 * x1854;
-Tbase TZ x1955 = 0x2;
-Tbase TZ x1956 = x1851 * x1955;
-Tbase TZ x1957 = x1839 * x1956;
-Tbase TZ x1958 = x1954 + x1957;
-Tbase TZ x1959 = x1953 + x1958;
-Tbase TZ x1960 = x1950 + x1959;
-Tbase TZ x1961 = x1949 + x1960;
-Tbase TZ x1962 = x1946 + x1961;
-Tbase TZ x1963 = x1945 + x1962;
-Tbase TZ x1964 = 0x13;
-Tbase TZ x1965 = x1964 * x1963;
-Tbase TZ x1966 = x1942 + x1965;
-Tbase TZ x1967 = x1935 + x1966;
-Tbase TZ x1968 = 0x1a;
-Tbase TZ x1969 = x1967 >> x1968;
-Tbase TZ x1970 = x1839 * x1878;
-Tbase TZ x1971 = x1842 * x1875;
-Tbase TZ x1972 = x1845 * x1872;
-Tbase TZ x1973 = x1848 * x1869;
-Tbase TZ x1974 = x1972 + x1973;
-Tbase TZ x1975 = x1971 + x1974;
-Tbase TZ x1976 = x1970 + x1975;
-Tbase TZ x1977 = x1821 * x1866;
-Tbase TZ x1978 = x1824 * x1863;
-Tbase TZ x1979 = x1827 * x1860;
-Tbase TZ x1980 = x1830 * x1857;
-Tbase TZ x1981 = x1833 * x1854;
-Tbase TZ x1982 = x1836 * x1851;
-Tbase TZ x1983 = x1981 + x1982;
-Tbase TZ x1984 = x1980 + x1983;
-Tbase TZ x1985 = x1979 + x1984;
-Tbase TZ x1986 = x1978 + x1985;
-Tbase TZ x1987 = x1977 + x1986;
-Tbase TZ x1988 = 0x13;
-Tbase TZ x1989 = x1988 * x1987;
-Tbase TZ x1990 = x1976 + x1989;
-Tbase TZ x1991 = x1969 + x1990;
-Tbase TZ x1992 = 0x19;
-Tbase TZ x1993 = x1991 >> x1992;
-Tbase TZ x1994 = x1836 * x1878;
-Tbase TZ x1995 = 0x2;
-Tbase TZ x1996 = x1875 * x1995;
-Tbase TZ x1997 = x1839 * x1996;
-Tbase TZ x1998 = x1842 * x1872;
-Tbase TZ x1999 = 0x2;
-Tbase TZ x2000 = x1869 * x1999;
-Tbase TZ x2001 = x1845 * x2000;
-Tbase TZ x2002 = x1848 * x1866;
-Tbase TZ x2003 = x2001 + x2002;
-Tbase TZ x2004 = x1998 + x2003;
-Tbase TZ x2005 = x1997 + x2004;
-Tbase TZ x2006 = x1994 + x2005;
-Tbase TZ x2007 = 0x2;
-Tbase TZ x2008 = x1863 * x2007;
-Tbase TZ x2009 = x1821 * x2008;
-Tbase TZ x2010 = x1824 * x1860;
-Tbase TZ x2011 = 0x2;
-Tbase TZ x2012 = x1857 * x2011;
-Tbase TZ x2013 = x1827 * x2012;
-Tbase TZ x2014 = x1830 * x1854;
-Tbase TZ x2015 = 0x2;
-Tbase TZ x2016 = x1851 * x2015;
-Tbase TZ x2017 = x1833 * x2016;
-Tbase TZ x2018 = x2014 + x2017;
-Tbase TZ x2019 = x2013 + x2018;
-Tbase TZ x2020 = x2010 + x2019;
-Tbase TZ x2021 = x2009 + x2020;
-Tbase TZ x2022 = 0x13;
-Tbase TZ x2023 = x2022 * x2021;
-Tbase TZ x2024 = x2006 + x2023;
-Tbase TZ x2025 = x1993 + x2024;
-Tbase TZ x2026 = 0x1a;
-Tbase TZ x2027 = x2025 >> x2026;
-Tbase TZ x2028 = x1833 * x1878;
-Tbase TZ x2029 = x1836 * x1875;
-Tbase TZ x2030 = x1839 * x1872;
-Tbase TZ x2031 = x1842 * x1869;
-Tbase TZ x2032 = x1845 * x1866;
-Tbase TZ x2033 = x1848 * x1863;
-Tbase TZ x2034 = x2032 + x2033;
-Tbase TZ x2035 = x2031 + x2034;
-Tbase TZ x2036 = x2030 + x2035;
-Tbase TZ x2037 = x2029 + x2036;
-Tbase TZ x2038 = x2028 + x2037;
-Tbase TZ x2039 = x1821 * x1860;
-Tbase TZ x2040 = x1824 * x1857;
-Tbase TZ x2041 = x1827 * x1854;
-Tbase TZ x2042 = x1830 * x1851;
-Tbase TZ x2043 = x2041 + x2042;
-Tbase TZ x2044 = x2040 + x2043;
-Tbase TZ x2045 = x2039 + x2044;
-Tbase TZ x2046 = 0x13;
-Tbase TZ x2047 = x2046 * x2045;
-Tbase TZ x2048 = x2038 + x2047;
-Tbase TZ x2049 = x2027 + x2048;
-Tbase TZ x2050 = 0x19;
-Tbase TZ x2051 = x2049 >> x2050;
-Tbase TZ x2052 = x1830 * x1878;
-Tbase TZ x2053 = 0x2;
-Tbase TZ x2054 = x1875 * x2053;
-Tbase TZ x2055 = x1833 * x2054;
-Tbase TZ x2056 = x1836 * x1872;
-Tbase TZ x2057 = 0x2;
-Tbase TZ x2058 = x1869 * x2057;
-Tbase TZ x2059 = x1839 * x2058;
-Tbase TZ x2060 = x1842 * x1866;
-Tbase TZ x2061 = 0x2;
-Tbase TZ x2062 = x1863 * x2061;
-Tbase TZ x2063 = x1845 * x2062;
-Tbase TZ x2064 = x1848 * x1860;
-Tbase TZ x2065 = x2063 + x2064;
-Tbase TZ x2066 = x2060 + x2065;
-Tbase TZ x2067 = x2059 + x2066;
-Tbase TZ x2068 = x2056 + x2067;
-Tbase TZ x2069 = x2055 + x2068;
-Tbase TZ x2070 = x2052 + x2069;
-Tbase TZ x2071 = 0x2;
-Tbase TZ x2072 = x1857 * x2071;
-Tbase TZ x2073 = x1821 * x2072;
-Tbase TZ x2074 = x1824 * x1854;
-Tbase TZ x2075 = 0x2;
-Tbase TZ x2076 = x1851 * x2075;
-Tbase TZ x2077 = x1827 * x2076;
-Tbase TZ x2078 = x2074 + x2077;
-Tbase TZ x2079 = x2073 + x2078;
-Tbase TZ x2080 = 0x13;
-Tbase TZ x2081 = x2080 * x2079;
-Tbase TZ x2082 = x2070 + x2081;
-Tbase TZ x2083 = x2051 + x2082;
-Tbase TZ x2084 = 0x1a;
-Tbase TZ x2085 = x2083 >> x2084;
-Tbase TZ x2086 = x1827 * x1878;
-Tbase TZ x2087 = x1830 * x1875;
-Tbase TZ x2088 = x1833 * x1872;
-Tbase TZ x2089 = x1836 * x1869;
-Tbase TZ x2090 = x1839 * x1866;
-Tbase TZ x2091 = x1842 * x1863;
-Tbase TZ x2092 = x1845 * x1860;
-Tbase TZ x2093 = x1848 * x1857;
-Tbase TZ x2094 = x2092 + x2093;
-Tbase TZ x2095 = x2091 + x2094;
-Tbase TZ x2096 = x2090 + x2095;
-Tbase TZ x2097 = x2089 + x2096;
-Tbase TZ x2098 = x2088 + x2097;
-Tbase TZ x2099 = x2087 + x2098;
-Tbase TZ x2100 = x2086 + x2099;
-Tbase TZ x2101 = x1821 * x1854;
-Tbase TZ x2102 = x1824 * x1851;
-Tbase TZ x2103 = x2101 + x2102;
-Tbase TZ x2104 = 0x13;
-Tbase TZ x2105 = x2104 * x2103;
-Tbase TZ x2106 = x2100 + x2105;
-Tbase TZ x2107 = x2085 + x2106;
-Tbase TZ x2108 = 0x19;
-Tbase TZ x2109 = x2107 >> x2108;
-Tbase TZ x2110 = x1824 * x1878;
-Tbase TZ x2111 = 0x2;
-Tbase TZ x2112 = x1875 * x2111;
-Tbase TZ x2113 = x1827 * x2112;
-Tbase TZ x2114 = x1830 * x1872;
-Tbase TZ x2115 = 0x2;
-Tbase TZ x2116 = x1869 * x2115;
-Tbase TZ x2117 = x1833 * x2116;
-Tbase TZ x2118 = x1836 * x1866;
-Tbase TZ x2119 = 0x2;
-Tbase TZ x2120 = x1863 * x2119;
-Tbase TZ x2121 = x1839 * x2120;
-Tbase TZ x2122 = x1842 * x1860;
-Tbase TZ x2123 = 0x2;
-Tbase TZ x2124 = x1857 * x2123;
-Tbase TZ x2125 = x1845 * x2124;
-Tbase TZ x2126 = x1848 * x1854;
-Tbase TZ x2127 = x2125 + x2126;
-Tbase TZ x2128 = x2122 + x2127;
-Tbase TZ x2129 = x2121 + x2128;
-Tbase TZ x2130 = x2118 + x2129;
-Tbase TZ x2131 = x2117 + x2130;
-Tbase TZ x2132 = x2114 + x2131;
-Tbase TZ x2133 = x2113 + x2132;
-Tbase TZ x2134 = x2110 + x2133;
-Tbase TZ x2135 = 0x2;
-Tbase TZ x2136 = x1851 * x2135;
-Tbase TZ x2137 = x1821 * x2136;
-Tbase TZ x2138 = 0x13;
-Tbase TZ x2139 = x2138 * x2137;
-Tbase TZ x2140 = x2134 + x2139;
-Tbase TZ x2141 = x2109 + x2140;
-Tbase TZ x2142 = 0x1a;
-Tbase TZ x2143 = x2141 >> x2142;
-Tbase TZ x2144 = x1821 * x1878;
-Tbase TZ x2145 = x1824 * x1875;
-Tbase TZ x2146 = x1827 * x1872;
-Tbase TZ x2147 = x1830 * x1869;
-Tbase TZ x2148 = x1833 * x1866;
-Tbase TZ x2149 = x1836 * x1863;
-Tbase TZ x2150 = x1839 * x1860;
-Tbase TZ x2151 = x1842 * x1857;
-Tbase TZ x2152 = x1845 * x1854;
-Tbase TZ x2153 = x1848 * x1851;
-Tbase TZ x2154 = x2152 + x2153;
-Tbase TZ x2155 = x2151 + x2154;
-Tbase TZ x2156 = x2150 + x2155;
-Tbase TZ x2157 = x2149 + x2156;
-Tbase TZ x2158 = x2148 + x2157;
-Tbase TZ x2159 = x2147 + x2158;
-Tbase TZ x2160 = x2146 + x2159;
-Tbase TZ x2161 = x2145 + x2160;
-Tbase TZ x2162 = x2144 + x2161;
-Tbase TZ x2163 = x2143 + x2162;
-Tbase TZ x2164 = 0x19;
-Tbase TZ x2165 = x2163 >> x2164;
-Tbase TZ x2166 = 0x13;
-Tbase TZ x2167 = x2166 * x2165;
-Tbase TZ x2168 = 0x3ffffff;
-Tbase TZ x2169 = x1909 & x2168;
-Tbase TZ x2170 = x2167 + x2169;
-Tbase TZ x2171 = 0x1a;
-Tbase TZ x2172 = x2170 >> x2171;
-Tbase TZ x2173 = 0x1ffffff;
-Tbase TZ x2174 = x1933 & x2173;
-Tbase TZ x2175 = x2172 + x2174;
-Tbase TZ x2176 = 0x1ffffff;
-Tbase TZ x2177 = x2163 & x2176;
-Tbase TZ x2178 = 0x3ffffff;
-Tbase TZ x2179 = x2141 & x2178;
-Tbase TZ x2180 = 0x1ffffff;
-Tbase TZ x2181 = x2107 & x2180;
-Tbase TZ x2182 = 0x3ffffff;
-Tbase TZ x2183 = x2083 & x2182;
-Tbase TZ x2184 = 0x1ffffff;
-Tbase TZ x2185 = x2049 & x2184;
-Tbase TZ x2186 = 0x3ffffff;
-Tbase TZ x2187 = x2025 & x2186;
-Tbase TZ x2188 = 0x1ffffff;
-Tbase TZ x2189 = x1991 & x2188;
-Tbase TZ x2190 = 0x19;
-Tbase TZ x2191 = x2175 >> x2190;
-Tbase TZ x2192 = 0x3ffffff;
-Tbase TZ x2193 = x1967 & x2192;
-Tbase TZ x2194 = x2191 + x2193;
-Tbase TZ x2195 = 0x1ffffff;
-Tbase TZ x2196 = x2175 & x2195;
-Tbase TZ x2197 = 0x3ffffff;
-Tbase TZ x2198 = x2170 & x2197;
-Tbase TZ x2199 = x48 * x2198;
-Tbase TZ x2200 = 0x2;
-Tbase TZ x2201 = x2196 * x2200;
-Tbase TZ x2202 = x39 * x2201;
-Tbase TZ x2203 = x40 * x2194;
-Tbase TZ x2204 = 0x2;
-Tbase TZ x2205 = x2189 * x2204;
-Tbase TZ x2206 = x41 * x2205;
-Tbase TZ x2207 = x42 * x2187;
-Tbase TZ x2208 = 0x2;
-Tbase TZ x2209 = x2185 * x2208;
-Tbase TZ x2210 = x43 * x2209;
-Tbase TZ x2211 = x44 * x2183;
-Tbase TZ x2212 = 0x2;
-Tbase TZ x2213 = x2181 * x2212;
-Tbase TZ x2214 = x45 * x2213;
-Tbase TZ x2215 = x46 * x2179;
-Tbase TZ x2216 = 0x2;
-Tbase TZ x2217 = x2177 * x2216;
-Tbase TZ x2218 = x47 * x2217;
-Tbase TZ x2219 = x2215 + x2218;
-Tbase TZ x2220 = x2214 + x2219;
-Tbase TZ x2221 = x2211 + x2220;
-Tbase TZ x2222 = x2210 + x2221;
-Tbase TZ x2223 = x2207 + x2222;
-Tbase TZ x2224 = x2206 + x2223;
-Tbase TZ x2225 = x2203 + x2224;
-Tbase TZ x2226 = x2202 + x2225;
-Tbase TZ x2227 = 0x13;
-Tbase TZ x2228 = x2227 * x2226;
-Tbase TZ x2229 = x2199 + x2228;
-Tbase TZ x2230 = 0x1a;
-Tbase TZ x2231 = x2229 >> x2230;
-Tbase TZ x2232 = x47 * x2198;
-Tbase TZ x2233 = x48 * x2196;
-Tbase TZ x2234 = x2232 + x2233;
-Tbase TZ x2235 = x39 * x2194;
-Tbase TZ x2236 = x40 * x2189;
-Tbase TZ x2237 = x41 * x2187;
-Tbase TZ x2238 = x42 * x2185;
-Tbase TZ x2239 = x43 * x2183;
-Tbase TZ x2240 = x44 * x2181;
-Tbase TZ x2241 = x45 * x2179;
-Tbase TZ x2242 = x46 * x2177;
-Tbase TZ x2243 = x2241 + x2242;
-Tbase TZ x2244 = x2240 + x2243;
-Tbase TZ x2245 = x2239 + x2244;
-Tbase TZ x2246 = x2238 + x2245;
-Tbase TZ x2247 = x2237 + x2246;
-Tbase TZ x2248 = x2236 + x2247;
-Tbase TZ x2249 = x2235 + x2248;
-Tbase TZ x2250 = 0x13;
-Tbase TZ x2251 = x2250 * x2249;
-Tbase TZ x2252 = x2234 + x2251;
-Tbase TZ x2253 = x2231 + x2252;
-Tbase TZ x2254 = 0x19;
-Tbase TZ x2255 = x2253 >> x2254;
-Tbase TZ x2256 = x46 * x2198;
-Tbase TZ x2257 = 0x2;
-Tbase TZ x2258 = x2196 * x2257;
-Tbase TZ x2259 = x47 * x2258;
-Tbase TZ x2260 = x48 * x2194;
-Tbase TZ x2261 = x2259 + x2260;
-Tbase TZ x2262 = x2256 + x2261;
-Tbase TZ x2263 = 0x2;
-Tbase TZ x2264 = x2189 * x2263;
-Tbase TZ x2265 = x39 * x2264;
-Tbase TZ x2266 = x40 * x2187;
-Tbase TZ x2267 = 0x2;
-Tbase TZ x2268 = x2185 * x2267;
-Tbase TZ x2269 = x41 * x2268;
-Tbase TZ x2270 = x42 * x2183;
-Tbase TZ x2271 = 0x2;
-Tbase TZ x2272 = x2181 * x2271;
-Tbase TZ x2273 = x43 * x2272;
-Tbase TZ x2274 = x44 * x2179;
-Tbase TZ x2275 = 0x2;
-Tbase TZ x2276 = x2177 * x2275;
-Tbase TZ x2277 = x45 * x2276;
-Tbase TZ x2278 = x2274 + x2277;
-Tbase TZ x2279 = x2273 + x2278;
-Tbase TZ x2280 = x2270 + x2279;
-Tbase TZ x2281 = x2269 + x2280;
-Tbase TZ x2282 = x2266 + x2281;
-Tbase TZ x2283 = x2265 + x2282;
-Tbase TZ x2284 = 0x13;
-Tbase TZ x2285 = x2284 * x2283;
-Tbase TZ x2286 = x2262 + x2285;
-Tbase TZ x2287 = x2255 + x2286;
-Tbase TZ x2288 = 0x1a;
-Tbase TZ x2289 = x2287 >> x2288;
-Tbase TZ x2290 = x45 * x2198;
-Tbase TZ x2291 = x46 * x2196;
-Tbase TZ x2292 = x47 * x2194;
-Tbase TZ x2293 = x48 * x2189;
-Tbase TZ x2294 = x2292 + x2293;
-Tbase TZ x2295 = x2291 + x2294;
-Tbase TZ x2296 = x2290 + x2295;
-Tbase TZ x2297 = x39 * x2187;
-Tbase TZ x2298 = x40 * x2185;
-Tbase TZ x2299 = x41 * x2183;
-Tbase TZ x2300 = x42 * x2181;
-Tbase TZ x2301 = x43 * x2179;
-Tbase TZ x2302 = x44 * x2177;
-Tbase TZ x2303 = x2301 + x2302;
-Tbase TZ x2304 = x2300 + x2303;
-Tbase TZ x2305 = x2299 + x2304;
-Tbase TZ x2306 = x2298 + x2305;
-Tbase TZ x2307 = x2297 + x2306;
-Tbase TZ x2308 = 0x13;
-Tbase TZ x2309 = x2308 * x2307;
-Tbase TZ x2310 = x2296 + x2309;
-Tbase TZ x2311 = x2289 + x2310;
-Tbase TZ x2312 = 0x19;
-Tbase TZ x2313 = x2311 >> x2312;
-Tbase TZ x2314 = x44 * x2198;
-Tbase TZ x2315 = 0x2;
-Tbase TZ x2316 = x2196 * x2315;
-Tbase TZ x2317 = x45 * x2316;
-Tbase TZ x2318 = x46 * x2194;
-Tbase TZ x2319 = 0x2;
-Tbase TZ x2320 = x2189 * x2319;
-Tbase TZ x2321 = x47 * x2320;
-Tbase TZ x2322 = x48 * x2187;
-Tbase TZ x2323 = x2321 + x2322;
-Tbase TZ x2324 = x2318 + x2323;
-Tbase TZ x2325 = x2317 + x2324;
-Tbase TZ x2326 = x2314 + x2325;
-Tbase TZ x2327 = 0x2;
-Tbase TZ x2328 = x2185 * x2327;
-Tbase TZ x2329 = x39 * x2328;
-Tbase TZ x2330 = x40 * x2183;
-Tbase TZ x2331 = 0x2;
-Tbase TZ x2332 = x2181 * x2331;
-Tbase TZ x2333 = x41 * x2332;
-Tbase TZ x2334 = x42 * x2179;
-Tbase TZ x2335 = 0x2;
-Tbase TZ x2336 = x2177 * x2335;
-Tbase TZ x2337 = x43 * x2336;
-Tbase TZ x2338 = x2334 + x2337;
-Tbase TZ x2339 = x2333 + x2338;
-Tbase TZ x2340 = x2330 + x2339;
-Tbase TZ x2341 = x2329 + x2340;
-Tbase TZ x2342 = 0x13;
-Tbase TZ x2343 = x2342 * x2341;
-Tbase TZ x2344 = x2326 + x2343;
-Tbase TZ x2345 = x2313 + x2344;
-Tbase TZ x2346 = 0x1a;
-Tbase TZ x2347 = x2345 >> x2346;
-Tbase TZ x2348 = x43 * x2198;
-Tbase TZ x2349 = x44 * x2196;
-Tbase TZ x2350 = x45 * x2194;
-Tbase TZ x2351 = x46 * x2189;
-Tbase TZ x2352 = x47 * x2187;
-Tbase TZ x2353 = x48 * x2185;
-Tbase TZ x2354 = x2352 + x2353;
-Tbase TZ x2355 = x2351 + x2354;
-Tbase TZ x2356 = x2350 + x2355;
-Tbase TZ x2357 = x2349 + x2356;
-Tbase TZ x2358 = x2348 + x2357;
-Tbase TZ x2359 = x39 * x2183;
-Tbase TZ x2360 = x40 * x2181;
-Tbase TZ x2361 = x41 * x2179;
-Tbase TZ x2362 = x42 * x2177;
-Tbase TZ x2363 = x2361 + x2362;
-Tbase TZ x2364 = x2360 + x2363;
-Tbase TZ x2365 = x2359 + x2364;
-Tbase TZ x2366 = 0x13;
-Tbase TZ x2367 = x2366 * x2365;
-Tbase TZ x2368 = x2358 + x2367;
-Tbase TZ x2369 = x2347 + x2368;
-Tbase TZ x2370 = 0x19;
-Tbase TZ x2371 = x2369 >> x2370;
-Tbase TZ x2372 = x42 * x2198;
-Tbase TZ x2373 = 0x2;
-Tbase TZ x2374 = x2196 * x2373;
-Tbase TZ x2375 = x43 * x2374;
-Tbase TZ x2376 = x44 * x2194;
-Tbase TZ x2377 = 0x2;
-Tbase TZ x2378 = x2189 * x2377;
-Tbase TZ x2379 = x45 * x2378;
-Tbase TZ x2380 = x46 * x2187;
-Tbase TZ x2381 = 0x2;
-Tbase TZ x2382 = x2185 * x2381;
-Tbase TZ x2383 = x47 * x2382;
-Tbase TZ x2384 = x48 * x2183;
-Tbase TZ x2385 = x2383 + x2384;
-Tbase TZ x2386 = x2380 + x2385;
-Tbase TZ x2387 = x2379 + x2386;
-Tbase TZ x2388 = x2376 + x2387;
-Tbase TZ x2389 = x2375 + x2388;
-Tbase TZ x2390 = x2372 + x2389;
-Tbase TZ x2391 = 0x2;
-Tbase TZ x2392 = x2181 * x2391;
-Tbase TZ x2393 = x39 * x2392;
-Tbase TZ x2394 = x40 * x2179;
-Tbase TZ x2395 = 0x2;
-Tbase TZ x2396 = x2177 * x2395;
-Tbase TZ x2397 = x41 * x2396;
-Tbase TZ x2398 = x2394 + x2397;
-Tbase TZ x2399 = x2393 + x2398;
-Tbase TZ x2400 = 0x13;
-Tbase TZ x2401 = x2400 * x2399;
-Tbase TZ x2402 = x2390 + x2401;
-Tbase TZ x2403 = x2371 + x2402;
-Tbase TZ x2404 = 0x1a;
-Tbase TZ x2405 = x2403 >> x2404;
-Tbase TZ x2406 = x41 * x2198;
-Tbase TZ x2407 = x42 * x2196;
-Tbase TZ x2408 = x43 * x2194;
-Tbase TZ x2409 = x44 * x2189;
-Tbase TZ x2410 = x45 * x2187;
-Tbase TZ x2411 = x46 * x2185;
-Tbase TZ x2412 = x47 * x2183;
-Tbase TZ x2413 = x48 * x2181;
-Tbase TZ x2414 = x2412 + x2413;
-Tbase TZ x2415 = x2411 + x2414;
-Tbase TZ x2416 = x2410 + x2415;
-Tbase TZ x2417 = x2409 + x2416;
-Tbase TZ x2418 = x2408 + x2417;
-Tbase TZ x2419 = x2407 + x2418;
-Tbase TZ x2420 = x2406 + x2419;
-Tbase TZ x2421 = x39 * x2179;
-Tbase TZ x2422 = x40 * x2177;
-Tbase TZ x2423 = x2421 + x2422;
-Tbase TZ x2424 = 0x13;
-Tbase TZ x2425 = x2424 * x2423;
-Tbase TZ x2426 = x2420 + x2425;
-Tbase TZ x2427 = x2405 + x2426;
-Tbase TZ x2428 = 0x19;
-Tbase TZ x2429 = x2427 >> x2428;
-Tbase TZ x2430 = x40 * x2198;
-Tbase TZ x2431 = 0x2;
-Tbase TZ x2432 = x2196 * x2431;
-Tbase TZ x2433 = x41 * x2432;
-Tbase TZ x2434 = x42 * x2194;
-Tbase TZ x2435 = 0x2;
-Tbase TZ x2436 = x2189 * x2435;
-Tbase TZ x2437 = x43 * x2436;
-Tbase TZ x2438 = x44 * x2187;
-Tbase TZ x2439 = 0x2;
-Tbase TZ x2440 = x2185 * x2439;
-Tbase TZ x2441 = x45 * x2440;
-Tbase TZ x2442 = x46 * x2183;
-Tbase TZ x2443 = 0x2;
-Tbase TZ x2444 = x2181 * x2443;
-Tbase TZ x2445 = x47 * x2444;
-Tbase TZ x2446 = x48 * x2179;
-Tbase TZ x2447 = x2445 + x2446;
-Tbase TZ x2448 = x2442 + x2447;
-Tbase TZ x2449 = x2441 + x2448;
-Tbase TZ x2450 = x2438 + x2449;
-Tbase TZ x2451 = x2437 + x2450;
-Tbase TZ x2452 = x2434 + x2451;
-Tbase TZ x2453 = x2433 + x2452;
-Tbase TZ x2454 = x2430 + x2453;
-Tbase TZ x2455 = 0x2;
-Tbase TZ x2456 = x2177 * x2455;
-Tbase TZ x2457 = x39 * x2456;
-Tbase TZ x2458 = 0x13;
-Tbase TZ x2459 = x2458 * x2457;
-Tbase TZ x2460 = x2454 + x2459;
-Tbase TZ x2461 = x2429 + x2460;
-Tbase TZ x2462 = 0x1a;
-Tbase TZ x2463 = x2461 >> x2462;
-Tbase TZ x2464 = x39 * x2198;
-Tbase TZ x2465 = x40 * x2196;
-Tbase TZ x2466 = x41 * x2194;
-Tbase TZ x2467 = x42 * x2189;
-Tbase TZ x2468 = x43 * x2187;
-Tbase TZ x2469 = x44 * x2185;
-Tbase TZ x2470 = x45 * x2183;
-Tbase TZ x2471 = x46 * x2181;
-Tbase TZ x2472 = x47 * x2179;
-Tbase TZ x2473 = x48 * x2177;
-Tbase TZ x2474 = x2472 + x2473;
-Tbase TZ x2475 = x2471 + x2474;
-Tbase TZ x2476 = x2470 + x2475;
-Tbase TZ x2477 = x2469 + x2476;
-Tbase TZ x2478 = x2468 + x2477;
-Tbase TZ x2479 = x2467 + x2478;
-Tbase TZ x2480 = x2466 + x2479;
-Tbase TZ x2481 = x2465 + x2480;
-Tbase TZ x2482 = x2464 + x2481;
-Tbase TZ x2483 = x2463 + x2482;
-Tbase TZ x2484 = 0x19;
-Tbase TZ x2485 = x2483 >> x2484;
-Tbase TZ x2486 = 0x13;
-Tbase TZ x2487 = x2486 * x2485;
-Tbase TZ x2488 = 0x3ffffff;
-Tbase TZ x2489 = x2229 & x2488;
-Tbase TZ x2490 = x2487 + x2489;
-Tbase TZ x2491 = 0x1a;
-Tbase TZ x2492 = x2490 >> x2491;
-Tbase TZ x2493 = 0x1ffffff;
-Tbase TZ x2494 = x2253 & x2493;
-Tbase TZ x2495 = x2492 + x2494;
-Tbase TZ x2496 = 0x1ffffff;
-Tbase TZ x2497 = x2483 & x2496;
-Tbase TZ x2498 = 0x3ffffff;
-Tbase TZ x2499 = x2461 & x2498;
-Tbase TZ x2500 = 0x1ffffff;
-Tbase TZ x2501 = x2427 & x2500;
-Tbase TZ x2502 = 0x3ffffff;
-Tbase TZ x2503 = x2403 & x2502;
-Tbase TZ x2504 = 0x1ffffff;
-Tbase TZ x2505 = x2369 & x2504;
-Tbase TZ x2506 = 0x3ffffff;
-Tbase TZ x2507 = x2345 & x2506;
-Tbase TZ x2508 = 0x1ffffff;
-Tbase TZ x2509 = x2311 & x2508;
-Tbase TZ x2510 = 0x19;
-Tbase TZ x2511 = x2495 >> x2510;
-Tbase TZ x2512 = 0x3ffffff;
-Tbase TZ x2513 = x2287 & x2512;
-Tbase TZ x2514 = x2511 + x2513;
-Tbase TZ x2515 = 0x1ffffff;
-Tbase TZ x2516 = x2495 & x2515;
-Tbase TZ x2517 = 0x3ffffff;
-Tbase TZ x2518 = x2490 & x2517;
-Tbase TZ x2519 = x418 * x768;
-Tbase TZ x2520 = 0x2;
-Tbase TZ x2521 = x766 * x2520;
-Tbase TZ x2522 = x397 * x2521;
-Tbase TZ x2523 = x399 * x764;
-Tbase TZ x2524 = 0x2;
-Tbase TZ x2525 = x759 * x2524;
-Tbase TZ x2526 = x401 * x2525;
-Tbase TZ x2527 = x403 * x757;
-Tbase TZ x2528 = 0x2;
-Tbase TZ x2529 = x755 * x2528;
-Tbase TZ x2530 = x405 * x2529;
-Tbase TZ x2531 = x407 * x753;
-Tbase TZ x2532 = 0x2;
-Tbase TZ x2533 = x751 * x2532;
-Tbase TZ x2534 = x409 * x2533;
-Tbase TZ x2535 = x414 * x749;
-Tbase TZ x2536 = 0x2;
-Tbase TZ x2537 = x747 * x2536;
-Tbase TZ x2538 = x416 * x2537;
-Tbase TZ x2539 = x2535 + x2538;
-Tbase TZ x2540 = x2534 + x2539;
-Tbase TZ x2541 = x2531 + x2540;
-Tbase TZ x2542 = x2530 + x2541;
-Tbase TZ x2543 = x2527 + x2542;
-Tbase TZ x2544 = x2526 + x2543;
-Tbase TZ x2545 = x2523 + x2544;
-Tbase TZ x2546 = x2522 + x2545;
-Tbase TZ x2547 = 0x13;
-Tbase TZ x2548 = x2547 * x2546;
-Tbase TZ x2549 = x2519 + x2548;
-Tbase TZ x2550 = 0x1a;
-Tbase TZ x2551 = x2549 >> x2550;
-Tbase TZ x2552 = x416 * x768;
-Tbase TZ x2553 = x418 * x766;
-Tbase TZ x2554 = x2552 + x2553;
-Tbase TZ x2555 = x397 * x764;
-Tbase TZ x2556 = x399 * x759;
-Tbase TZ x2557 = x401 * x757;
-Tbase TZ x2558 = x403 * x755;
-Tbase TZ x2559 = x405 * x753;
-Tbase TZ x2560 = x407 * x751;
-Tbase TZ x2561 = x409 * x749;
-Tbase TZ x2562 = x414 * x747;
-Tbase TZ x2563 = x2561 + x2562;
-Tbase TZ x2564 = x2560 + x2563;
-Tbase TZ x2565 = x2559 + x2564;
-Tbase TZ x2566 = x2558 + x2565;
-Tbase TZ x2567 = x2557 + x2566;
-Tbase TZ x2568 = x2556 + x2567;
-Tbase TZ x2569 = x2555 + x2568;
-Tbase TZ x2570 = 0x13;
-Tbase TZ x2571 = x2570 * x2569;
-Tbase TZ x2572 = x2554 + x2571;
-Tbase TZ x2573 = x2551 + x2572;
-Tbase TZ x2574 = 0x19;
-Tbase TZ x2575 = x2573 >> x2574;
-Tbase TZ x2576 = x414 * x768;
-Tbase TZ x2577 = 0x2;
-Tbase TZ x2578 = x766 * x2577;
-Tbase TZ x2579 = x416 * x2578;
-Tbase TZ x2580 = x418 * x764;
-Tbase TZ x2581 = x2579 + x2580;
-Tbase TZ x2582 = x2576 + x2581;
-Tbase TZ x2583 = 0x2;
-Tbase TZ x2584 = x759 * x2583;
-Tbase TZ x2585 = x397 * x2584;
-Tbase TZ x2586 = x399 * x757;
-Tbase TZ x2587 = 0x2;
-Tbase TZ x2588 = x755 * x2587;
-Tbase TZ x2589 = x401 * x2588;
-Tbase TZ x2590 = x403 * x753;
-Tbase TZ x2591 = 0x2;
-Tbase TZ x2592 = x751 * x2591;
-Tbase TZ x2593 = x405 * x2592;
-Tbase TZ x2594 = x407 * x749;
-Tbase TZ x2595 = 0x2;
-Tbase TZ x2596 = x747 * x2595;
-Tbase TZ x2597 = x409 * x2596;
-Tbase TZ x2598 = x2594 + x2597;
-Tbase TZ x2599 = x2593 + x2598;
-Tbase TZ x2600 = x2590 + x2599;
-Tbase TZ x2601 = x2589 + x2600;
-Tbase TZ x2602 = x2586 + x2601;
-Tbase TZ x2603 = x2585 + x2602;
-Tbase TZ x2604 = 0x13;
-Tbase TZ x2605 = x2604 * x2603;
-Tbase TZ x2606 = x2582 + x2605;
-Tbase TZ x2607 = x2575 + x2606;
-Tbase TZ x2608 = 0x1a;
-Tbase TZ x2609 = x2607 >> x2608;
-Tbase TZ x2610 = x409 * x768;
-Tbase TZ x2611 = x414 * x766;
-Tbase TZ x2612 = x416 * x764;
-Tbase TZ x2613 = x418 * x759;
-Tbase TZ x2614 = x2612 + x2613;
-Tbase TZ x2615 = x2611 + x2614;
-Tbase TZ x2616 = x2610 + x2615;
-Tbase TZ x2617 = x397 * x757;
-Tbase TZ x2618 = x399 * x755;
-Tbase TZ x2619 = x401 * x753;
-Tbase TZ x2620 = x403 * x751;
-Tbase TZ x2621 = x405 * x749;
-Tbase TZ x2622 = x407 * x747;
-Tbase TZ x2623 = x2621 + x2622;
-Tbase TZ x2624 = x2620 + x2623;
-Tbase TZ x2625 = x2619 + x2624;
-Tbase TZ x2626 = x2618 + x2625;
-Tbase TZ x2627 = x2617 + x2626;
-Tbase TZ x2628 = 0x13;
-Tbase TZ x2629 = x2628 * x2627;
-Tbase TZ x2630 = x2616 + x2629;
-Tbase TZ x2631 = x2609 + x2630;
-Tbase TZ x2632 = 0x19;
-Tbase TZ x2633 = x2631 >> x2632;
-Tbase TZ x2634 = x407 * x768;
-Tbase TZ x2635 = 0x2;
-Tbase TZ x2636 = x766 * x2635;
-Tbase TZ x2637 = x409 * x2636;
-Tbase TZ x2638 = x414 * x764;
-Tbase TZ x2639 = 0x2;
-Tbase TZ x2640 = x759 * x2639;
-Tbase TZ x2641 = x416 * x2640;
-Tbase TZ x2642 = x418 * x757;
-Tbase TZ x2643 = x2641 + x2642;
-Tbase TZ x2644 = x2638 + x2643;
-Tbase TZ x2645 = x2637 + x2644;
-Tbase TZ x2646 = x2634 + x2645;
-Tbase TZ x2647 = 0x2;
-Tbase TZ x2648 = x755 * x2647;
-Tbase TZ x2649 = x397 * x2648;
-Tbase TZ x2650 = x399 * x753;
-Tbase TZ x2651 = 0x2;
-Tbase TZ x2652 = x751 * x2651;
-Tbase TZ x2653 = x401 * x2652;
-Tbase TZ x2654 = x403 * x749;
-Tbase TZ x2655 = 0x2;
-Tbase TZ x2656 = x747 * x2655;
-Tbase TZ x2657 = x405 * x2656;
-Tbase TZ x2658 = x2654 + x2657;
-Tbase TZ x2659 = x2653 + x2658;
-Tbase TZ x2660 = x2650 + x2659;
-Tbase TZ x2661 = x2649 + x2660;
-Tbase TZ x2662 = 0x13;
-Tbase TZ x2663 = x2662 * x2661;
-Tbase TZ x2664 = x2646 + x2663;
-Tbase TZ x2665 = x2633 + x2664;
-Tbase TZ x2666 = 0x1a;
-Tbase TZ x2667 = x2665 >> x2666;
-Tbase TZ x2668 = x405 * x768;
-Tbase TZ x2669 = x407 * x766;
-Tbase TZ x2670 = x409 * x764;
-Tbase TZ x2671 = x414 * x759;
-Tbase TZ x2672 = x416 * x757;
-Tbase TZ x2673 = x418 * x755;
-Tbase TZ x2674 = x2672 + x2673;
-Tbase TZ x2675 = x2671 + x2674;
-Tbase TZ x2676 = x2670 + x2675;
-Tbase TZ x2677 = x2669 + x2676;
-Tbase TZ x2678 = x2668 + x2677;
-Tbase TZ x2679 = x397 * x753;
-Tbase TZ x2680 = x399 * x751;
-Tbase TZ x2681 = x401 * x749;
-Tbase TZ x2682 = x403 * x747;
-Tbase TZ x2683 = x2681 + x2682;
-Tbase TZ x2684 = x2680 + x2683;
-Tbase TZ x2685 = x2679 + x2684;
-Tbase TZ x2686 = 0x13;
-Tbase TZ x2687 = x2686 * x2685;
-Tbase TZ x2688 = x2678 + x2687;
-Tbase TZ x2689 = x2667 + x2688;
-Tbase TZ x2690 = 0x19;
-Tbase TZ x2691 = x2689 >> x2690;
-Tbase TZ x2692 = x403 * x768;
-Tbase TZ x2693 = 0x2;
-Tbase TZ x2694 = x766 * x2693;
-Tbase TZ x2695 = x405 * x2694;
-Tbase TZ x2696 = x407 * x764;
-Tbase TZ x2697 = 0x2;
-Tbase TZ x2698 = x759 * x2697;
-Tbase TZ x2699 = x409 * x2698;
-Tbase TZ x2700 = x414 * x757;
-Tbase TZ x2701 = 0x2;
-Tbase TZ x2702 = x755 * x2701;
-Tbase TZ x2703 = x416 * x2702;
-Tbase TZ x2704 = x418 * x753;
-Tbase TZ x2705 = x2703 + x2704;
-Tbase TZ x2706 = x2700 + x2705;
-Tbase TZ x2707 = x2699 + x2706;
-Tbase TZ x2708 = x2696 + x2707;
-Tbase TZ x2709 = x2695 + x2708;
-Tbase TZ x2710 = x2692 + x2709;
-Tbase TZ x2711 = 0x2;
-Tbase TZ x2712 = x751 * x2711;
-Tbase TZ x2713 = x397 * x2712;
-Tbase TZ x2714 = x399 * x749;
-Tbase TZ x2715 = 0x2;
-Tbase TZ x2716 = x747 * x2715;
-Tbase TZ x2717 = x401 * x2716;
-Tbase TZ x2718 = x2714 + x2717;
-Tbase TZ x2719 = x2713 + x2718;
-Tbase TZ x2720 = 0x13;
-Tbase TZ x2721 = x2720 * x2719;
-Tbase TZ x2722 = x2710 + x2721;
-Tbase TZ x2723 = x2691 + x2722;
-Tbase TZ x2724 = 0x1a;
-Tbase TZ x2725 = x2723 >> x2724;
-Tbase TZ x2726 = x401 * x768;
-Tbase TZ x2727 = x403 * x766;
-Tbase TZ x2728 = x405 * x764;
-Tbase TZ x2729 = x407 * x759;
-Tbase TZ x2730 = x409 * x757;
-Tbase TZ x2731 = x414 * x755;
-Tbase TZ x2732 = x416 * x753;
-Tbase TZ x2733 = x418 * x751;
-Tbase TZ x2734 = x2732 + x2733;
-Tbase TZ x2735 = x2731 + x2734;
-Tbase TZ x2736 = x2730 + x2735;
-Tbase TZ x2737 = x2729 + x2736;
-Tbase TZ x2738 = x2728 + x2737;
-Tbase TZ x2739 = x2727 + x2738;
-Tbase TZ x2740 = x2726 + x2739;
-Tbase TZ x2741 = x397 * x749;
-Tbase TZ x2742 = x399 * x747;
-Tbase TZ x2743 = x2741 + x2742;
-Tbase TZ x2744 = 0x13;
-Tbase TZ x2745 = x2744 * x2743;
-Tbase TZ x2746 = x2740 + x2745;
-Tbase TZ x2747 = x2725 + x2746;
-Tbase TZ x2748 = 0x19;
-Tbase TZ x2749 = x2747 >> x2748;
-Tbase TZ x2750 = x399 * x768;
-Tbase TZ x2751 = 0x2;
-Tbase TZ x2752 = x766 * x2751;
-Tbase TZ x2753 = x401 * x2752;
-Tbase TZ x2754 = x403 * x764;
-Tbase TZ x2755 = 0x2;
-Tbase TZ x2756 = x759 * x2755;
-Tbase TZ x2757 = x405 * x2756;
-Tbase TZ x2758 = x407 * x757;
-Tbase TZ x2759 = 0x2;
-Tbase TZ x2760 = x755 * x2759;
-Tbase TZ x2761 = x409 * x2760;
-Tbase TZ x2762 = x414 * x753;
-Tbase TZ x2763 = 0x2;
-Tbase TZ x2764 = x751 * x2763;
-Tbase TZ x2765 = x416 * x2764;
-Tbase TZ x2766 = x418 * x749;
-Tbase TZ x2767 = x2765 + x2766;
-Tbase TZ x2768 = x2762 + x2767;
-Tbase TZ x2769 = x2761 + x2768;
-Tbase TZ x2770 = x2758 + x2769;
-Tbase TZ x2771 = x2757 + x2770;
-Tbase TZ x2772 = x2754 + x2771;
-Tbase TZ x2773 = x2753 + x2772;
-Tbase TZ x2774 = x2750 + x2773;
-Tbase TZ x2775 = 0x2;
-Tbase TZ x2776 = x747 * x2775;
-Tbase TZ x2777 = x397 * x2776;
-Tbase TZ x2778 = 0x13;
-Tbase TZ x2779 = x2778 * x2777;
-Tbase TZ x2780 = x2774 + x2779;
-Tbase TZ x2781 = x2749 + x2780;
-Tbase TZ x2782 = 0x1a;
-Tbase TZ x2783 = x2781 >> x2782;
-Tbase TZ x2784 = x397 * x768;
-Tbase TZ x2785 = x399 * x766;
-Tbase TZ x2786 = x401 * x764;
-Tbase TZ x2787 = x403 * x759;
-Tbase TZ x2788 = x405 * x757;
-Tbase TZ x2789 = x407 * x755;
-Tbase TZ x2790 = x409 * x753;
-Tbase TZ x2791 = x414 * x751;
-Tbase TZ x2792 = x416 * x749;
-Tbase TZ x2793 = x418 * x747;
-Tbase TZ x2794 = x2792 + x2793;
-Tbase TZ x2795 = x2791 + x2794;
-Tbase TZ x2796 = x2790 + x2795;
-Tbase TZ x2797 = x2789 + x2796;
-Tbase TZ x2798 = x2788 + x2797;
-Tbase TZ x2799 = x2787 + x2798;
-Tbase TZ x2800 = x2786 + x2799;
-Tbase TZ x2801 = x2785 + x2800;
-Tbase TZ x2802 = x2784 + x2801;
-Tbase TZ x2803 = x2783 + x2802;
-Tbase TZ x2804 = 0x19;
-Tbase TZ x2805 = x2803 >> x2804;
-Tbase TZ x2806 = 0x13;
-Tbase TZ x2807 = x2806 * x2805;
-Tbase TZ x2808 = 0x3ffffff;
-Tbase TZ x2809 = x2549 & x2808;
-Tbase TZ x2810 = x2807 + x2809;
-Tbase TZ x2811 = 0x1a;
-Tbase TZ x2812 = x2810 >> x2811;
-Tbase TZ x2813 = 0x1ffffff;
-Tbase TZ x2814 = x2573 & x2813;
-Tbase TZ x2815 = x2812 + x2814;
-Tbase TZ x2816 = 0x1ffffff;
-Tbase TZ x2817 = x2803 & x2816;
-Tbase TZ x2818 = 0x3ffffff;
-Tbase TZ x2819 = x2781 & x2818;
-Tbase TZ x2820 = 0x1ffffff;
-Tbase TZ x2821 = x2747 & x2820;
-Tbase TZ x2822 = 0x3ffffff;
-Tbase TZ x2823 = x2723 & x2822;
-Tbase TZ x2824 = 0x1ffffff;
-Tbase TZ x2825 = x2689 & x2824;
-Tbase TZ x2826 = 0x3ffffff;
-Tbase TZ x2827 = x2665 & x2826;
-Tbase TZ x2828 = 0x1ffffff;
-Tbase TZ x2829 = x2631 & x2828;
-Tbase TZ x2830 = 0x19;
-Tbase TZ x2831 = x2815 >> x2830;
-Tbase TZ x2832 = 0x3ffffff;
-Tbase TZ x2833 = x2607 & x2832;
-Tbase TZ x2834 = x2831 + x2833;
-Tbase TZ x2835 = 0x1ffffff;
-Tbase TZ x2836 = x2815 & x2835;
-Tbase TZ x2837 = 0x3ffffff;
-Tbase TZ x2838 = x2810 & x2837;
-Tbase TZ x2839 = x38 * x798;
-Tbase TZ x2840 = 0x2;
-Tbase TZ x2841 = x795 * x2840;
-Tbase TZ x2842 = x29 * x2841;
-Tbase TZ x2843 = x30 * x792;
-Tbase TZ x2844 = 0x2;
-Tbase TZ x2845 = x789 * x2844;
-Tbase TZ x2846 = x31 * x2845;
-Tbase TZ x2847 = x32 * x786;
-Tbase TZ x2848 = 0x2;
-Tbase TZ x2849 = x783 * x2848;
-Tbase TZ x2850 = x33 * x2849;
-Tbase TZ x2851 = x34 * x780;
-Tbase TZ x2852 = 0x2;
-Tbase TZ x2853 = x777 * x2852;
-Tbase TZ x2854 = x35 * x2853;
-Tbase TZ x2855 = x36 * x774;
-Tbase TZ x2856 = 0x2;
-Tbase TZ x2857 = x771 * x2856;
-Tbase TZ x2858 = x37 * x2857;
-Tbase TZ x2859 = x2855 + x2858;
-Tbase TZ x2860 = x2854 + x2859;
-Tbase TZ x2861 = x2851 + x2860;
-Tbase TZ x2862 = x2850 + x2861;
-Tbase TZ x2863 = x2847 + x2862;
-Tbase TZ x2864 = x2846 + x2863;
-Tbase TZ x2865 = x2843 + x2864;
-Tbase TZ x2866 = x2842 + x2865;
-Tbase TZ x2867 = 0x13;
-Tbase TZ x2868 = x2867 * x2866;
-Tbase TZ x2869 = x2839 + x2868;
-Tbase TZ x2870 = 0x1a;
-Tbase TZ x2871 = x2869 >> x2870;
-Tbase TZ x2872 = x37 * x798;
-Tbase TZ x2873 = x38 * x795;
-Tbase TZ x2874 = x2872 + x2873;
-Tbase TZ x2875 = x29 * x792;
-Tbase TZ x2876 = x30 * x789;
-Tbase TZ x2877 = x31 * x786;
-Tbase TZ x2878 = x32 * x783;
-Tbase TZ x2879 = x33 * x780;
-Tbase TZ x2880 = x34 * x777;
-Tbase TZ x2881 = x35 * x774;
-Tbase TZ x2882 = x36 * x771;
-Tbase TZ x2883 = x2881 + x2882;
-Tbase TZ x2884 = x2880 + x2883;
-Tbase TZ x2885 = x2879 + x2884;
-Tbase TZ x2886 = x2878 + x2885;
-Tbase TZ x2887 = x2877 + x2886;
-Tbase TZ x2888 = x2876 + x2887;
-Tbase TZ x2889 = x2875 + x2888;
-Tbase TZ x2890 = 0x13;
-Tbase TZ x2891 = x2890 * x2889;
-Tbase TZ x2892 = x2874 + x2891;
-Tbase TZ x2893 = x2871 + x2892;
-Tbase TZ x2894 = 0x19;
-Tbase TZ x2895 = x2893 >> x2894;
-Tbase TZ x2896 = x36 * x798;
-Tbase TZ x2897 = 0x2;
-Tbase TZ x2898 = x795 * x2897;
-Tbase TZ x2899 = x37 * x2898;
-Tbase TZ x2900 = x38 * x792;
-Tbase TZ x2901 = x2899 + x2900;
-Tbase TZ x2902 = x2896 + x2901;
-Tbase TZ x2903 = 0x2;
-Tbase TZ x2904 = x789 * x2903;
-Tbase TZ x2905 = x29 * x2904;
-Tbase TZ x2906 = x30 * x786;
-Tbase TZ x2907 = 0x2;
-Tbase TZ x2908 = x783 * x2907;
-Tbase TZ x2909 = x31 * x2908;
-Tbase TZ x2910 = x32 * x780;
-Tbase TZ x2911 = 0x2;
-Tbase TZ x2912 = x777 * x2911;
-Tbase TZ x2913 = x33 * x2912;
-Tbase TZ x2914 = x34 * x774;
-Tbase TZ x2915 = 0x2;
-Tbase TZ x2916 = x771 * x2915;
-Tbase TZ x2917 = x35 * x2916;
-Tbase TZ x2918 = x2914 + x2917;
-Tbase TZ x2919 = x2913 + x2918;
-Tbase TZ x2920 = x2910 + x2919;
-Tbase TZ x2921 = x2909 + x2920;
-Tbase TZ x2922 = x2906 + x2921;
-Tbase TZ x2923 = x2905 + x2922;
-Tbase TZ x2924 = 0x13;
-Tbase TZ x2925 = x2924 * x2923;
-Tbase TZ x2926 = x2902 + x2925;
-Tbase TZ x2927 = x2895 + x2926;
-Tbase TZ x2928 = 0x1a;
-Tbase TZ x2929 = x2927 >> x2928;
-Tbase TZ x2930 = x35 * x798;
-Tbase TZ x2931 = x36 * x795;
-Tbase TZ x2932 = x37 * x792;
-Tbase TZ x2933 = x38 * x789;
-Tbase TZ x2934 = x2932 + x2933;
-Tbase TZ x2935 = x2931 + x2934;
-Tbase TZ x2936 = x2930 + x2935;
-Tbase TZ x2937 = x29 * x786;
-Tbase TZ x2938 = x30 * x783;
-Tbase TZ x2939 = x31 * x780;
-Tbase TZ x2940 = x32 * x777;
-Tbase TZ x2941 = x33 * x774;
-Tbase TZ x2942 = x34 * x771;
-Tbase TZ x2943 = x2941 + x2942;
-Tbase TZ x2944 = x2940 + x2943;
-Tbase TZ x2945 = x2939 + x2944;
-Tbase TZ x2946 = x2938 + x2945;
-Tbase TZ x2947 = x2937 + x2946;
-Tbase TZ x2948 = 0x13;
-Tbase TZ x2949 = x2948 * x2947;
-Tbase TZ x2950 = x2936 + x2949;
-Tbase TZ x2951 = x2929 + x2950;
-Tbase TZ x2952 = 0x19;
-Tbase TZ x2953 = x2951 >> x2952;
-Tbase TZ x2954 = x34 * x798;
-Tbase TZ x2955 = 0x2;
-Tbase TZ x2956 = x795 * x2955;
-Tbase TZ x2957 = x35 * x2956;
-Tbase TZ x2958 = x36 * x792;
-Tbase TZ x2959 = 0x2;
-Tbase TZ x2960 = x789 * x2959;
-Tbase TZ x2961 = x37 * x2960;
-Tbase TZ x2962 = x38 * x786;
-Tbase TZ x2963 = x2961 + x2962;
-Tbase TZ x2964 = x2958 + x2963;
-Tbase TZ x2965 = x2957 + x2964;
-Tbase TZ x2966 = x2954 + x2965;
-Tbase TZ x2967 = 0x2;
-Tbase TZ x2968 = x783 * x2967;
-Tbase TZ x2969 = x29 * x2968;
-Tbase TZ x2970 = x30 * x780;
-Tbase TZ x2971 = 0x2;
-Tbase TZ x2972 = x777 * x2971;
-Tbase TZ x2973 = x31 * x2972;
-Tbase TZ x2974 = x32 * x774;
-Tbase TZ x2975 = 0x2;
-Tbase TZ x2976 = x771 * x2975;
-Tbase TZ x2977 = x33 * x2976;
-Tbase TZ x2978 = x2974 + x2977;
-Tbase TZ x2979 = x2973 + x2978;
-Tbase TZ x2980 = x2970 + x2979;
-Tbase TZ x2981 = x2969 + x2980;
-Tbase TZ x2982 = 0x13;
-Tbase TZ x2983 = x2982 * x2981;
-Tbase TZ x2984 = x2966 + x2983;
-Tbase TZ x2985 = x2953 + x2984;
-Tbase TZ x2986 = 0x1a;
-Tbase TZ x2987 = x2985 >> x2986;
-Tbase TZ x2988 = x33 * x798;
-Tbase TZ x2989 = x34 * x795;
-Tbase TZ x2990 = x35 * x792;
-Tbase TZ x2991 = x36 * x789;
-Tbase TZ x2992 = x37 * x786;
-Tbase TZ x2993 = x38 * x783;
-Tbase TZ x2994 = x2992 + x2993;
-Tbase TZ x2995 = x2991 + x2994;
-Tbase TZ x2996 = x2990 + x2995;
-Tbase TZ x2997 = x2989 + x2996;
-Tbase TZ x2998 = x2988 + x2997;
-Tbase TZ x2999 = x29 * x780;
-Tbase TZ x3000 = x30 * x777;
-Tbase TZ x3001 = x31 * x774;
-Tbase TZ x3002 = x32 * x771;
-Tbase TZ x3003 = x3001 + x3002;
-Tbase TZ x3004 = x3000 + x3003;
-Tbase TZ x3005 = x2999 + x3004;
-Tbase TZ x3006 = 0x13;
-Tbase TZ x3007 = x3006 * x3005;
-Tbase TZ x3008 = x2998 + x3007;
-Tbase TZ x3009 = x2987 + x3008;
-Tbase TZ x3010 = 0x19;
-Tbase TZ x3011 = x3009 >> x3010;
-Tbase TZ x3012 = x32 * x798;
-Tbase TZ x3013 = 0x2;
-Tbase TZ x3014 = x795 * x3013;
-Tbase TZ x3015 = x33 * x3014;
-Tbase TZ x3016 = x34 * x792;
-Tbase TZ x3017 = 0x2;
-Tbase TZ x3018 = x789 * x3017;
-Tbase TZ x3019 = x35 * x3018;
-Tbase TZ x3020 = x36 * x786;
-Tbase TZ x3021 = 0x2;
-Tbase TZ x3022 = x783 * x3021;
-Tbase TZ x3023 = x37 * x3022;
-Tbase TZ x3024 = x38 * x780;
-Tbase TZ x3025 = x3023 + x3024;
-Tbase TZ x3026 = x3020 + x3025;
-Tbase TZ x3027 = x3019 + x3026;
-Tbase TZ x3028 = x3016 + x3027;
-Tbase TZ x3029 = x3015 + x3028;
-Tbase TZ x3030 = x3012 + x3029;
-Tbase TZ x3031 = 0x2;
-Tbase TZ x3032 = x777 * x3031;
-Tbase TZ x3033 = x29 * x3032;
-Tbase TZ x3034 = x30 * x774;
-Tbase TZ x3035 = 0x2;
-Tbase TZ x3036 = x771 * x3035;
-Tbase TZ x3037 = x31 * x3036;
-Tbase TZ x3038 = x3034 + x3037;
-Tbase TZ x3039 = x3033 + x3038;
-Tbase TZ x3040 = 0x13;
-Tbase TZ x3041 = x3040 * x3039;
-Tbase TZ x3042 = x3030 + x3041;
-Tbase TZ x3043 = x3011 + x3042;
-Tbase TZ x3044 = 0x1a;
-Tbase TZ x3045 = x3043 >> x3044;
-Tbase TZ x3046 = x31 * x798;
-Tbase TZ x3047 = x32 * x795;
-Tbase TZ x3048 = x33 * x792;
-Tbase TZ x3049 = x34 * x789;
-Tbase TZ x3050 = x35 * x786;
-Tbase TZ x3051 = x36 * x783;
-Tbase TZ x3052 = x37 * x780;
-Tbase TZ x3053 = x38 * x777;
-Tbase TZ x3054 = x3052 + x3053;
-Tbase TZ x3055 = x3051 + x3054;
-Tbase TZ x3056 = x3050 + x3055;
-Tbase TZ x3057 = x3049 + x3056;
-Tbase TZ x3058 = x3048 + x3057;
-Tbase TZ x3059 = x3047 + x3058;
-Tbase TZ x3060 = x3046 + x3059;
-Tbase TZ x3061 = x29 * x774;
-Tbase TZ x3062 = x30 * x771;
-Tbase TZ x3063 = x3061 + x3062;
-Tbase TZ x3064 = 0x13;
-Tbase TZ x3065 = x3064 * x3063;
-Tbase TZ x3066 = x3060 + x3065;
-Tbase TZ x3067 = x3045 + x3066;
-Tbase TZ x3068 = 0x19;
-Tbase TZ x3069 = x3067 >> x3068;
-Tbase TZ x3070 = x30 * x798;
-Tbase TZ x3071 = 0x2;
-Tbase TZ x3072 = x795 * x3071;
-Tbase TZ x3073 = x31 * x3072;
-Tbase TZ x3074 = x32 * x792;
-Tbase TZ x3075 = 0x2;
-Tbase TZ x3076 = x789 * x3075;
-Tbase TZ x3077 = x33 * x3076;
-Tbase TZ x3078 = x34 * x786;
-Tbase TZ x3079 = 0x2;
-Tbase TZ x3080 = x783 * x3079;
-Tbase TZ x3081 = x35 * x3080;
-Tbase TZ x3082 = x36 * x780;
-Tbase TZ x3083 = 0x2;
-Tbase TZ x3084 = x777 * x3083;
-Tbase TZ x3085 = x37 * x3084;
-Tbase TZ x3086 = x38 * x774;
-Tbase TZ x3087 = x3085 + x3086;
-Tbase TZ x3088 = x3082 + x3087;
-Tbase TZ x3089 = x3081 + x3088;
-Tbase TZ x3090 = x3078 + x3089;
-Tbase TZ x3091 = x3077 + x3090;
-Tbase TZ x3092 = x3074 + x3091;
-Tbase TZ x3093 = x3073 + x3092;
-Tbase TZ x3094 = x3070 + x3093;
-Tbase TZ x3095 = 0x2;
-Tbase TZ x3096 = x771 * x3095;
-Tbase TZ x3097 = x29 * x3096;
-Tbase TZ x3098 = 0x13;
-Tbase TZ x3099 = x3098 * x3097;
-Tbase TZ x3100 = x3094 + x3099;
-Tbase TZ x3101 = x3069 + x3100;
-Tbase TZ x3102 = 0x1a;
-Tbase TZ x3103 = x3101 >> x3102;
-Tbase TZ x3104 = x29 * x798;
-Tbase TZ x3105 = x30 * x795;
-Tbase TZ x3106 = x31 * x792;
-Tbase TZ x3107 = x32 * x789;
-Tbase TZ x3108 = x33 * x786;
-Tbase TZ x3109 = x34 * x783;
-Tbase TZ x3110 = x35 * x780;
-Tbase TZ x3111 = x36 * x777;
-Tbase TZ x3112 = x37 * x774;
-Tbase TZ x3113 = x38 * x771;
-Tbase TZ x3114 = x3112 + x3113;
-Tbase TZ x3115 = x3111 + x3114;
-Tbase TZ x3116 = x3110 + x3115;
-Tbase TZ x3117 = x3109 + x3116;
-Tbase TZ x3118 = x3108 + x3117;
-Tbase TZ x3119 = x3107 + x3118;
-Tbase TZ x3120 = x3106 + x3119;
-Tbase TZ x3121 = x3105 + x3120;
-Tbase TZ x3122 = x3104 + x3121;
-Tbase TZ x3123 = x3103 + x3122;
-Tbase TZ x3124 = 0x19;
-Tbase TZ x3125 = x3123 >> x3124;
-Tbase TZ x3126 = 0x13;
-Tbase TZ x3127 = x3126 * x3125;
-Tbase TZ x3128 = 0x3ffffff;
-Tbase TZ x3129 = x2869 & x3128;
-Tbase TZ x3130 = x3127 + x3129;
-Tbase TZ x3131 = 0x1a;
-Tbase TZ x3132 = x3130 >> x3131;
-Tbase TZ x3133 = 0x1ffffff;
-Tbase TZ x3134 = x2893 & x3133;
-Tbase TZ x3135 = x3132 + x3134;
-Tbase TZ x3136 = 0x1ffffff;
-Tbase TZ x3137 = x3123 & x3136;
-Tbase TZ x3138 = 0x3ffffff;
-Tbase TZ x3139 = x3101 & x3138;
-Tbase TZ x3140 = 0x1ffffff;
-Tbase TZ x3141 = x3067 & x3140;
-Tbase TZ x3142 = 0x3ffffff;
-Tbase TZ x3143 = x3043 & x3142;
-Tbase TZ x3144 = 0x1ffffff;
-Tbase TZ x3145 = x3009 & x3144;
-Tbase TZ x3146 = 0x3ffffff;
-Tbase TZ x3147 = x2985 & x3146;
-Tbase TZ x3148 = 0x1ffffff;
-Tbase TZ x3149 = x2951 & x3148;
-Tbase TZ x3150 = 0x19;
-Tbase TZ x3151 = x3135 >> x3150;
-Tbase TZ x3152 = 0x3ffffff;
-Tbase TZ x3153 = x2927 & x3152;
-Tbase TZ x3154 = x3151 + x3153;
-Tbase TZ x3155 = 0x1ffffff;
-Tbase TZ x3156 = x3135 & x3155;
-Tbase TZ x3157 = 0x3ffffff;
-Tbase TZ x3158 = x3130 & x3157;
-Tbase TZ x3159 = x397 + x3137;
-Tbase TZ x3160 = x399 + x3139;
-Tbase TZ x3161 = x401 + x3141;
-Tbase TZ x3162 = x403 + x3143;
-Tbase TZ x3163 = x405 + x3145;
-Tbase TZ x3164 = x407 + x3147;
-Tbase TZ x3165 = x409 + x3149;
-Tbase TZ x3166 = x414 + x3154;
-Tbase TZ x3167 = x416 + x3156;
-Tbase TZ x3168 = x418 + x3158;
-Tbase TZ x3169 = x798 * x3168;
-Tbase TZ x3170 = 0x2;
-Tbase TZ x3171 = x3167 * x3170;
-Tbase TZ x3172 = x771 * x3171;
-Tbase TZ x3173 = x774 * x3166;
-Tbase TZ x3174 = 0x2;
-Tbase TZ x3175 = x3165 * x3174;
-Tbase TZ x3176 = x777 * x3175;
-Tbase TZ x3177 = x780 * x3164;
-Tbase TZ x3178 = 0x2;
-Tbase TZ x3179 = x3163 * x3178;
-Tbase TZ x3180 = x783 * x3179;
-Tbase TZ x3181 = x786 * x3162;
-Tbase TZ x3182 = 0x2;
-Tbase TZ x3183 = x3161 * x3182;
-Tbase TZ x3184 = x789 * x3183;
-Tbase TZ x3185 = x792 * x3160;
-Tbase TZ x3186 = 0x2;
-Tbase TZ x3187 = x3159 * x3186;
-Tbase TZ x3188 = x795 * x3187;
-Tbase TZ x3189 = x3185 + x3188;
-Tbase TZ x3190 = x3184 + x3189;
-Tbase TZ x3191 = x3181 + x3190;
-Tbase TZ x3192 = x3180 + x3191;
-Tbase TZ x3193 = x3177 + x3192;
-Tbase TZ x3194 = x3176 + x3193;
-Tbase TZ x3195 = x3173 + x3194;
-Tbase TZ x3196 = x3172 + x3195;
-Tbase TZ x3197 = 0x13;
-Tbase TZ x3198 = x3197 * x3196;
-Tbase TZ x3199 = x3169 + x3198;
-Tbase TZ x3200 = 0x1a;
-Tbase TZ x3201 = x3199 >> x3200;
-Tbase TZ x3202 = x795 * x3168;
-Tbase TZ x3203 = x798 * x3167;
-Tbase TZ x3204 = x3202 + x3203;
-Tbase TZ x3205 = x771 * x3166;
-Tbase TZ x3206 = x774 * x3165;
-Tbase TZ x3207 = x777 * x3164;
-Tbase TZ x3208 = x780 * x3163;
-Tbase TZ x3209 = x783 * x3162;
-Tbase TZ x3210 = x786 * x3161;
-Tbase TZ x3211 = x789 * x3160;
-Tbase TZ x3212 = x792 * x3159;
-Tbase TZ x3213 = x3211 + x3212;
-Tbase TZ x3214 = x3210 + x3213;
-Tbase TZ x3215 = x3209 + x3214;
-Tbase TZ x3216 = x3208 + x3215;
-Tbase TZ x3217 = x3207 + x3216;
-Tbase TZ x3218 = x3206 + x3217;
-Tbase TZ x3219 = x3205 + x3218;
-Tbase TZ x3220 = 0x13;
-Tbase TZ x3221 = x3220 * x3219;
-Tbase TZ x3222 = x3204 + x3221;
-Tbase TZ x3223 = x3201 + x3222;
-Tbase TZ x3224 = 0x19;
-Tbase TZ x3225 = x3223 >> x3224;
-Tbase TZ x3226 = x792 * x3168;
-Tbase TZ x3227 = 0x2;
-Tbase TZ x3228 = x3167 * x3227;
-Tbase TZ x3229 = x795 * x3228;
-Tbase TZ x3230 = x798 * x3166;
-Tbase TZ x3231 = x3229 + x3230;
-Tbase TZ x3232 = x3226 + x3231;
-Tbase TZ x3233 = 0x2;
-Tbase TZ x3234 = x3165 * x3233;
-Tbase TZ x3235 = x771 * x3234;
-Tbase TZ x3236 = x774 * x3164;
-Tbase TZ x3237 = 0x2;
-Tbase TZ x3238 = x3163 * x3237;
-Tbase TZ x3239 = x777 * x3238;
-Tbase TZ x3240 = x780 * x3162;
-Tbase TZ x3241 = 0x2;
-Tbase TZ x3242 = x3161 * x3241;
-Tbase TZ x3243 = x783 * x3242;
-Tbase TZ x3244 = x786 * x3160;
-Tbase TZ x3245 = 0x2;
-Tbase TZ x3246 = x3159 * x3245;
-Tbase TZ x3247 = x789 * x3246;
-Tbase TZ x3248 = x3244 + x3247;
-Tbase TZ x3249 = x3243 + x3248;
-Tbase TZ x3250 = x3240 + x3249;
-Tbase TZ x3251 = x3239 + x3250;
-Tbase TZ x3252 = x3236 + x3251;
-Tbase TZ x3253 = x3235 + x3252;
-Tbase TZ x3254 = 0x13;
-Tbase TZ x3255 = x3254 * x3253;
-Tbase TZ x3256 = x3232 + x3255;
-Tbase TZ x3257 = x3225 + x3256;
-Tbase TZ x3258 = 0x1a;
-Tbase TZ x3259 = x3257 >> x3258;
-Tbase TZ x3260 = x789 * x3168;
-Tbase TZ x3261 = x792 * x3167;
-Tbase TZ x3262 = x795 * x3166;
-Tbase TZ x3263 = x798 * x3165;
-Tbase TZ x3264 = x3262 + x3263;
-Tbase TZ x3265 = x3261 + x3264;
-Tbase TZ x3266 = x3260 + x3265;
-Tbase TZ x3267 = x771 * x3164;
-Tbase TZ x3268 = x774 * x3163;
-Tbase TZ x3269 = x777 * x3162;
-Tbase TZ x3270 = x780 * x3161;
-Tbase TZ x3271 = x783 * x3160;
-Tbase TZ x3272 = x786 * x3159;
-Tbase TZ x3273 = x3271 + x3272;
-Tbase TZ x3274 = x3270 + x3273;
-Tbase TZ x3275 = x3269 + x3274;
-Tbase TZ x3276 = x3268 + x3275;
-Tbase TZ x3277 = x3267 + x3276;
-Tbase TZ x3278 = 0x13;
-Tbase TZ x3279 = x3278 * x3277;
-Tbase TZ x3280 = x3266 + x3279;
-Tbase TZ x3281 = x3259 + x3280;
-Tbase TZ x3282 = 0x19;
-Tbase TZ x3283 = x3281 >> x3282;
-Tbase TZ x3284 = x786 * x3168;
-Tbase TZ x3285 = 0x2;
-Tbase TZ x3286 = x3167 * x3285;
-Tbase TZ x3287 = x789 * x3286;
-Tbase TZ x3288 = x792 * x3166;
-Tbase TZ x3289 = 0x2;
-Tbase TZ x3290 = x3165 * x3289;
-Tbase TZ x3291 = x795 * x3290;
-Tbase TZ x3292 = x798 * x3164;
-Tbase TZ x3293 = x3291 + x3292;
-Tbase TZ x3294 = x3288 + x3293;
-Tbase TZ x3295 = x3287 + x3294;
-Tbase TZ x3296 = x3284 + x3295;
-Tbase TZ x3297 = 0x2;
-Tbase TZ x3298 = x3163 * x3297;
-Tbase TZ x3299 = x771 * x3298;
-Tbase TZ x3300 = x774 * x3162;
-Tbase TZ x3301 = 0x2;
-Tbase TZ x3302 = x3161 * x3301;
-Tbase TZ x3303 = x777 * x3302;
-Tbase TZ x3304 = x780 * x3160;
-Tbase TZ x3305 = 0x2;
-Tbase TZ x3306 = x3159 * x3305;
-Tbase TZ x3307 = x783 * x3306;
-Tbase TZ x3308 = x3304 + x3307;
-Tbase TZ x3309 = x3303 + x3308;
-Tbase TZ x3310 = x3300 + x3309;
-Tbase TZ x3311 = x3299 + x3310;
-Tbase TZ x3312 = 0x13;
-Tbase TZ x3313 = x3312 * x3311;
-Tbase TZ x3314 = x3296 + x3313;
-Tbase TZ x3315 = x3283 + x3314;
-Tbase TZ x3316 = 0x1a;
-Tbase TZ x3317 = x3315 >> x3316;
-Tbase TZ x3318 = x783 * x3168;
-Tbase TZ x3319 = x786 * x3167;
-Tbase TZ x3320 = x789 * x3166;
-Tbase TZ x3321 = x792 * x3165;
-Tbase TZ x3322 = x795 * x3164;
-Tbase TZ x3323 = x798 * x3163;
-Tbase TZ x3324 = x3322 + x3323;
-Tbase TZ x3325 = x3321 + x3324;
-Tbase TZ x3326 = x3320 + x3325;
-Tbase TZ x3327 = x3319 + x3326;
-Tbase TZ x3328 = x3318 + x3327;
-Tbase TZ x3329 = x771 * x3162;
-Tbase TZ x3330 = x774 * x3161;
-Tbase TZ x3331 = x777 * x3160;
-Tbase TZ x3332 = x780 * x3159;
-Tbase TZ x3333 = x3331 + x3332;
-Tbase TZ x3334 = x3330 + x3333;
-Tbase TZ x3335 = x3329 + x3334;
-Tbase TZ x3336 = 0x13;
-Tbase TZ x3337 = x3336 * x3335;
-Tbase TZ x3338 = x3328 + x3337;
-Tbase TZ x3339 = x3317 + x3338;
-Tbase TZ x3340 = 0x19;
-Tbase TZ x3341 = x3339 >> x3340;
-Tbase TZ x3342 = x780 * x3168;
-Tbase TZ x3343 = 0x2;
-Tbase TZ x3344 = x3167 * x3343;
-Tbase TZ x3345 = x783 * x3344;
-Tbase TZ x3346 = x786 * x3166;
-Tbase TZ x3347 = 0x2;
-Tbase TZ x3348 = x3165 * x3347;
-Tbase TZ x3349 = x789 * x3348;
-Tbase TZ x3350 = x792 * x3164;
-Tbase TZ x3351 = 0x2;
-Tbase TZ x3352 = x3163 * x3351;
-Tbase TZ x3353 = x795 * x3352;
-Tbase TZ x3354 = x798 * x3162;
-Tbase TZ x3355 = x3353 + x3354;
-Tbase TZ x3356 = x3350 + x3355;
-Tbase TZ x3357 = x3349 + x3356;
-Tbase TZ x3358 = x3346 + x3357;
-Tbase TZ x3359 = x3345 + x3358;
-Tbase TZ x3360 = x3342 + x3359;
-Tbase TZ x3361 = 0x2;
-Tbase TZ x3362 = x3161 * x3361;
-Tbase TZ x3363 = x771 * x3362;
-Tbase TZ x3364 = x774 * x3160;
-Tbase TZ x3365 = 0x2;
-Tbase TZ x3366 = x3159 * x3365;
-Tbase TZ x3367 = x777 * x3366;
-Tbase TZ x3368 = x3364 + x3367;
-Tbase TZ x3369 = x3363 + x3368;
-Tbase TZ x3370 = 0x13;
-Tbase TZ x3371 = x3370 * x3369;
-Tbase TZ x3372 = x3360 + x3371;
-Tbase TZ x3373 = x3341 + x3372;
-Tbase TZ x3374 = 0x1a;
-Tbase TZ x3375 = x3373 >> x3374;
-Tbase TZ x3376 = x777 * x3168;
-Tbase TZ x3377 = x780 * x3167;
-Tbase TZ x3378 = x783 * x3166;
-Tbase TZ x3379 = x786 * x3165;
-Tbase TZ x3380 = x789 * x3164;
-Tbase TZ x3381 = x792 * x3163;
-Tbase TZ x3382 = x795 * x3162;
-Tbase TZ x3383 = x798 * x3161;
-Tbase TZ x3384 = x3382 + x3383;
-Tbase TZ x3385 = x3381 + x3384;
-Tbase TZ x3386 = x3380 + x3385;
-Tbase TZ x3387 = x3379 + x3386;
-Tbase TZ x3388 = x3378 + x3387;
-Tbase TZ x3389 = x3377 + x3388;
-Tbase TZ x3390 = x3376 + x3389;
-Tbase TZ x3391 = x771 * x3160;
-Tbase TZ x3392 = x774 * x3159;
-Tbase TZ x3393 = x3391 + x3392;
-Tbase TZ x3394 = 0x13;
-Tbase TZ x3395 = x3394 * x3393;
-Tbase TZ x3396 = x3390 + x3395;
-Tbase TZ x3397 = x3375 + x3396;
-Tbase TZ x3398 = 0x19;
-Tbase TZ x3399 = x3397 >> x3398;
-Tbase TZ x3400 = x774 * x3168;
-Tbase TZ x3401 = 0x2;
-Tbase TZ x3402 = x3167 * x3401;
-Tbase TZ x3403 = x777 * x3402;
-Tbase TZ x3404 = x780 * x3166;
-Tbase TZ x3405 = 0x2;
-Tbase TZ x3406 = x3165 * x3405;
-Tbase TZ x3407 = x783 * x3406;
-Tbase TZ x3408 = x786 * x3164;
-Tbase TZ x3409 = 0x2;
-Tbase TZ x3410 = x3163 * x3409;
-Tbase TZ x3411 = x789 * x3410;
-Tbase TZ x3412 = x792 * x3162;
-Tbase TZ x3413 = 0x2;
-Tbase TZ x3414 = x3161 * x3413;
-Tbase TZ x3415 = x795 * x3414;
-Tbase TZ x3416 = x798 * x3160;
-Tbase TZ x3417 = x3415 + x3416;
-Tbase TZ x3418 = x3412 + x3417;
-Tbase TZ x3419 = x3411 + x3418;
-Tbase TZ x3420 = x3408 + x3419;
-Tbase TZ x3421 = x3407 + x3420;
-Tbase TZ x3422 = x3404 + x3421;
-Tbase TZ x3423 = x3403 + x3422;
-Tbase TZ x3424 = x3400 + x3423;
-Tbase TZ x3425 = 0x2;
-Tbase TZ x3426 = x3159 * x3425;
-Tbase TZ x3427 = x771 * x3426;
-Tbase TZ x3428 = 0x13;
-Tbase TZ x3429 = x3428 * x3427;
-Tbase TZ x3430 = x3424 + x3429;
-Tbase TZ x3431 = x3399 + x3430;
-Tbase TZ x3432 = 0x1a;
-Tbase TZ x3433 = x3431 >> x3432;
-Tbase TZ x3434 = x771 * x3168;
-Tbase TZ x3435 = x774 * x3167;
-Tbase TZ x3436 = x777 * x3166;
-Tbase TZ x3437 = x780 * x3165;
-Tbase TZ x3438 = x783 * x3164;
-Tbase TZ x3439 = x786 * x3163;
-Tbase TZ x3440 = x789 * x3162;
-Tbase TZ x3441 = x792 * x3161;
-Tbase TZ x3442 = x795 * x3160;
-Tbase TZ x3443 = x798 * x3159;
-Tbase TZ x3444 = x3442 + x3443;
-Tbase TZ x3445 = x3441 + x3444;
-Tbase TZ x3446 = x3440 + x3445;
-Tbase TZ x3447 = x3439 + x3446;
-Tbase TZ x3448 = x3438 + x3447;
-Tbase TZ x3449 = x3437 + x3448;
-Tbase TZ x3450 = x3436 + x3449;
-Tbase TZ x3451 = x3435 + x3450;
-Tbase TZ x3452 = x3434 + x3451;
-Tbase TZ x3453 = x3433 + x3452;
-Tbase TZ x3454 = 0x19;
-Tbase TZ x3455 = x3453 >> x3454;
-Tbase TZ x3456 = 0x13;
-Tbase TZ x3457 = x3456 * x3455;
-Tbase TZ x3458 = 0x3ffffff;
-Tbase TZ x3459 = x3199 & x3458;
-Tbase TZ x3460 = x3457 + x3459;
-Tbase TZ x3461 = 0x1a;
-Tbase TZ x3462 = x3460 >> x3461;
-Tbase TZ x3463 = 0x1ffffff;
-Tbase TZ x3464 = x3223 & x3463;
-Tbase TZ x3465 = x3462 + x3464;
-Tbase TZ x3466 = 0x1ffffff;
-Tbase TZ x3467 = x3453 & x3466;
-Tbase TZ x3468 = 0x3ffffff;
-Tbase TZ x3469 = x3431 & x3468;
-Tbase TZ x3470 = 0x1ffffff;
-Tbase TZ x3471 = x3397 & x3470;
-Tbase TZ x3472 = 0x3ffffff;
-Tbase TZ x3473 = x3373 & x3472;
-Tbase TZ x3474 = 0x1ffffff;
-Tbase TZ x3475 = x3339 & x3474;
-Tbase TZ x3476 = 0x3ffffff;
-Tbase TZ x3477 = x3315 & x3476;
-Tbase TZ x3478 = 0x1ffffff;
-Tbase TZ x3479 = x3281 & x3478;
-Tbase TZ x3480 = 0x19;
-Tbase TZ x3481 = x3465 >> x3480;
-Tbase TZ x3482 = 0x3ffffff;
-Tbase TZ x3483 = x3257 & x3482;
-Tbase TZ x3484 = x3481 + x3483;
-Tbase TZ x3485 = 0x1ffffff;
-Tbase TZ x3486 = x3465 & x3485;
-Tbase TZ x3487 = 0x3ffffff;
-Tbase TZ x3488 = x3460 & x3487;
-(Return x2817, Return x2819, Return x2821, Return x2823,
-Return x2825, Return x2827, Return x2829, Return x2834,
-Return x2836, Return x2838,
-(Return x3467, Return x3469, Return x3471, Return x3473,
-Return x3475, Return x3477, Return x3479, Return x3484,
-Return x3486, Return x3488),
-(Return x1797, Return x1799, Return x1801, Return x1803,
-Return x1805, Return x1807, Return x1809, Return x1814,
-Return x1816, Return x1818),
-(Return x2497, Return x2499, Return x2501, Return x2503,
-Return x2505, Return x2507, Return x2509, Return x2514,
-Return x2516, Return x2518))
- : forall var : base_type -> Type,
- expr base_type op
- (TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> ...)
-
-Argument scope is [function_scope]
diff --git a/src/Specific/GF25519Reflective/Reified/LadderStepDisplay.v b/src/Specific/GF25519Reflective/Reified/LadderStepDisplay.v
deleted file mode 100644
index 9e395e4eb..000000000
--- a/src/Specific/GF25519Reflective/Reified/LadderStepDisplay.v
+++ /dev/null
@@ -1,4 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.Reified.LadderStep.
-Require Export Crypto.Reflection.Z.CNotations.
-
-Print rladderstepW.
diff --git a/src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.log b/src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.log
deleted file mode 100644
index 7fdca4c10..000000000
--- a/src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.log
+++ /dev/null
@@ -1,3459 +0,0 @@
-rladderstepW =
-fun var : base_type -> Type =>
-_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ x29 x30 x31 x32
- x33 x34 x35 x36 x37 x38 x39 x40 x41 x42 x43 x44 x45 x46 x47 x48 x49 x50 x51
- x52 x53 x54 x55 x56 x57 x58 x59 x60 x61 x62 x63 x64 x65 x66 x67 x68 x69 x70
- x71 x72 x73 x74 x75 x76 x77 x78 x79 x80 x81 x82 x83 x84 x85 x86 x87
- x88 : var TZ,
-Tbase TZ x89 = x49 + x59;
-Tbase TZ x90 = x50 + x60;
-Tbase TZ x91 = x51 + x61;
-Tbase TZ x92 = x52 + x62;
-Tbase TZ x93 = x53 + x63;
-Tbase TZ x94 = x54 + x64;
-Tbase TZ x95 = x55 + x65;
-Tbase TZ x96 = x56 + x66;
-Tbase TZ x97 = x57 + x67;
-Tbase TZ x98 = x58 + x68;
-Tbase TZ x99 = x98 * x98;
-Tbase TZ x100 = 0x2;
-Tbase TZ x101 = x97 * x100;
-Tbase TZ x102 = x89 * x101;
-Tbase TZ x103 = x90 * x96;
-Tbase TZ x104 = 0x2;
-Tbase TZ x105 = x95 * x104;
-Tbase TZ x106 = x91 * x105;
-Tbase TZ x107 = x92 * x94;
-Tbase TZ x108 = 0x2;
-Tbase TZ x109 = x93 * x108;
-Tbase TZ x110 = x93 * x109;
-Tbase TZ x111 = x94 * x92;
-Tbase TZ x112 = 0x2;
-Tbase TZ x113 = x91 * x112;
-Tbase TZ x114 = x95 * x113;
-Tbase TZ x115 = x96 * x90;
-Tbase TZ x116 = 0x2;
-Tbase TZ x117 = x89 * x116;
-Tbase TZ x118 = x97 * x117;
-Tbase TZ x119 = x115 + x118;
-Tbase TZ x120 = x114 + x119;
-Tbase TZ x121 = x111 + x120;
-Tbase TZ x122 = x110 + x121;
-Tbase TZ x123 = x107 + x122;
-Tbase TZ x124 = x106 + x123;
-Tbase TZ x125 = x103 + x124;
-Tbase TZ x126 = x102 + x125;
-Tbase TZ x127 = 0x13;
-Tbase TZ x128 = x127 * x126;
-Tbase TZ x129 = x99 + x128;
-Tbase TZ x130 = 0x1a;
-Tbase TZ x131 = x129 >>> x130;
-Tbase TZ x132 = x97 * x98;
-Tbase TZ x133 = x98 * x97;
-Tbase TZ x134 = x132 + x133;
-Tbase TZ x135 = x89 * x96;
-Tbase TZ x136 = x90 * x95;
-Tbase TZ x137 = x91 * x94;
-Tbase TZ x138 = x92 * x93;
-Tbase TZ x139 = x93 * x92;
-Tbase TZ x140 = x94 * x91;
-Tbase TZ x141 = x95 * x90;
-Tbase TZ x142 = x96 * x89;
-Tbase TZ x143 = x141 + x142;
-Tbase TZ x144 = x140 + x143;
-Tbase TZ x145 = x139 + x144;
-Tbase TZ x146 = x138 + x145;
-Tbase TZ x147 = x137 + x146;
-Tbase TZ x148 = x136 + x147;
-Tbase TZ x149 = x135 + x148;
-Tbase TZ x150 = 0x13;
-Tbase TZ x151 = x150 * x149;
-Tbase TZ x152 = x134 + x151;
-Tbase TZ x153 = x131 + x152;
-Tbase TZ x154 = 0x19;
-Tbase TZ x155 = x153 >>> x154;
-Tbase TZ x156 = x96 * x98;
-Tbase TZ x157 = 0x2;
-Tbase TZ x158 = x97 * x157;
-Tbase TZ x159 = x97 * x158;
-Tbase TZ x160 = x98 * x96;
-Tbase TZ x161 = x159 + x160;
-Tbase TZ x162 = x156 + x161;
-Tbase TZ x163 = 0x2;
-Tbase TZ x164 = x95 * x163;
-Tbase TZ x165 = x89 * x164;
-Tbase TZ x166 = x90 * x94;
-Tbase TZ x167 = 0x2;
-Tbase TZ x168 = x93 * x167;
-Tbase TZ x169 = x91 * x168;
-Tbase TZ x170 = x92 * x92;
-Tbase TZ x171 = 0x2;
-Tbase TZ x172 = x91 * x171;
-Tbase TZ x173 = x93 * x172;
-Tbase TZ x174 = x94 * x90;
-Tbase TZ x175 = 0x2;
-Tbase TZ x176 = x89 * x175;
-Tbase TZ x177 = x95 * x176;
-Tbase TZ x178 = x174 + x177;
-Tbase TZ x179 = x173 + x178;
-Tbase TZ x180 = x170 + x179;
-Tbase TZ x181 = x169 + x180;
-Tbase TZ x182 = x166 + x181;
-Tbase TZ x183 = x165 + x182;
-Tbase TZ x184 = 0x13;
-Tbase TZ x185 = x184 * x183;
-Tbase TZ x186 = x162 + x185;
-Tbase TZ x187 = x155 + x186;
-Tbase TZ x188 = 0x1a;
-Tbase TZ x189 = x187 >>> x188;
-Tbase TZ x190 = x95 * x98;
-Tbase TZ x191 = x96 * x97;
-Tbase TZ x192 = x97 * x96;
-Tbase TZ x193 = x98 * x95;
-Tbase TZ x194 = x192 + x193;
-Tbase TZ x195 = x191 + x194;
-Tbase TZ x196 = x190 + x195;
-Tbase TZ x197 = x89 * x94;
-Tbase TZ x198 = x90 * x93;
-Tbase TZ x199 = x91 * x92;
-Tbase TZ x200 = x92 * x91;
-Tbase TZ x201 = x93 * x90;
-Tbase TZ x202 = x94 * x89;
-Tbase TZ x203 = x201 + x202;
-Tbase TZ x204 = x200 + x203;
-Tbase TZ x205 = x199 + x204;
-Tbase TZ x206 = x198 + x205;
-Tbase TZ x207 = x197 + x206;
-Tbase TZ x208 = 0x13;
-Tbase TZ x209 = x208 * x207;
-Tbase TZ x210 = x196 + x209;
-Tbase TZ x211 = x189 + x210;
-Tbase TZ x212 = 0x19;
-Tbase TZ x213 = x211 >>> x212;
-Tbase TZ x214 = x94 * x98;
-Tbase TZ x215 = 0x2;
-Tbase TZ x216 = x97 * x215;
-Tbase TZ x217 = x95 * x216;
-Tbase TZ x218 = x96 * x96;
-Tbase TZ x219 = 0x2;
-Tbase TZ x220 = x95 * x219;
-Tbase TZ x221 = x97 * x220;
-Tbase TZ x222 = x98 * x94;
-Tbase TZ x223 = x221 + x222;
-Tbase TZ x224 = x218 + x223;
-Tbase TZ x225 = x217 + x224;
-Tbase TZ x226 = x214 + x225;
-Tbase TZ x227 = 0x2;
-Tbase TZ x228 = x93 * x227;
-Tbase TZ x229 = x89 * x228;
-Tbase TZ x230 = x90 * x92;
-Tbase TZ x231 = 0x2;
-Tbase TZ x232 = x91 * x231;
-Tbase TZ x233 = x91 * x232;
-Tbase TZ x234 = x92 * x90;
-Tbase TZ x235 = 0x2;
-Tbase TZ x236 = x89 * x235;
-Tbase TZ x237 = x93 * x236;
-Tbase TZ x238 = x234 + x237;
-Tbase TZ x239 = x233 + x238;
-Tbase TZ x240 = x230 + x239;
-Tbase TZ x241 = x229 + x240;
-Tbase TZ x242 = 0x13;
-Tbase TZ x243 = x242 * x241;
-Tbase TZ x244 = x226 + x243;
-Tbase TZ x245 = x213 + x244;
-Tbase TZ x246 = 0x1a;
-Tbase TZ x247 = x245 >>> x246;
-Tbase TZ x248 = x93 * x98;
-Tbase TZ x249 = x94 * x97;
-Tbase TZ x250 = x95 * x96;
-Tbase TZ x251 = x96 * x95;
-Tbase TZ x252 = x97 * x94;
-Tbase TZ x253 = x98 * x93;
-Tbase TZ x254 = x252 + x253;
-Tbase TZ x255 = x251 + x254;
-Tbase TZ x256 = x250 + x255;
-Tbase TZ x257 = x249 + x256;
-Tbase TZ x258 = x248 + x257;
-Tbase TZ x259 = x89 * x92;
-Tbase TZ x260 = x90 * x91;
-Tbase TZ x261 = x91 * x90;
-Tbase TZ x262 = x92 * x89;
-Tbase TZ x263 = x261 + x262;
-Tbase TZ x264 = x260 + x263;
-Tbase TZ x265 = x259 + x264;
-Tbase TZ x266 = 0x13;
-Tbase TZ x267 = x266 * x265;
-Tbase TZ x268 = x258 + x267;
-Tbase TZ x269 = x247 + x268;
-Tbase TZ x270 = 0x19;
-Tbase TZ x271 = x269 >>> x270;
-Tbase TZ x272 = x92 * x98;
-Tbase TZ x273 = 0x2;
-Tbase TZ x274 = x97 * x273;
-Tbase TZ x275 = x93 * x274;
-Tbase TZ x276 = x94 * x96;
-Tbase TZ x277 = 0x2;
-Tbase TZ x278 = x95 * x277;
-Tbase TZ x279 = x95 * x278;
-Tbase TZ x280 = x96 * x94;
-Tbase TZ x281 = 0x2;
-Tbase TZ x282 = x93 * x281;
-Tbase TZ x283 = x97 * x282;
-Tbase TZ x284 = x98 * x92;
-Tbase TZ x285 = x283 + x284;
-Tbase TZ x286 = x280 + x285;
-Tbase TZ x287 = x279 + x286;
-Tbase TZ x288 = x276 + x287;
-Tbase TZ x289 = x275 + x288;
-Tbase TZ x290 = x272 + x289;
-Tbase TZ x291 = 0x2;
-Tbase TZ x292 = x91 * x291;
-Tbase TZ x293 = x89 * x292;
-Tbase TZ x294 = x90 * x90;
-Tbase TZ x295 = 0x2;
-Tbase TZ x296 = x89 * x295;
-Tbase TZ x297 = x91 * x296;
-Tbase TZ x298 = x294 + x297;
-Tbase TZ x299 = x293 + x298;
-Tbase TZ x300 = 0x13;
-Tbase TZ x301 = x300 * x299;
-Tbase TZ x302 = x290 + x301;
-Tbase TZ x303 = x271 + x302;
-Tbase TZ x304 = 0x1a;
-Tbase TZ x305 = x303 >>> x304;
-Tbase TZ x306 = x91 * x98;
-Tbase TZ x307 = x92 * x97;
-Tbase TZ x308 = x93 * x96;
-Tbase TZ x309 = x94 * x95;
-Tbase TZ x310 = x95 * x94;
-Tbase TZ x311 = x96 * x93;
-Tbase TZ x312 = x97 * x92;
-Tbase TZ x313 = x98 * x91;
-Tbase TZ x314 = x312 + x313;
-Tbase TZ x315 = x311 + x314;
-Tbase TZ x316 = x310 + x315;
-Tbase TZ x317 = x309 + x316;
-Tbase TZ x318 = x308 + x317;
-Tbase TZ x319 = x307 + x318;
-Tbase TZ x320 = x306 + x319;
-Tbase TZ x321 = x89 * x90;
-Tbase TZ x322 = x90 * x89;
-Tbase TZ x323 = x321 + x322;
-Tbase TZ x324 = 0x13;
-Tbase TZ x325 = x324 * x323;
-Tbase TZ x326 = x320 + x325;
-Tbase TZ x327 = x305 + x326;
-Tbase TZ x328 = 0x19;
-Tbase TZ x329 = x327 >>> x328;
-Tbase TZ x330 = x90 * x98;
-Tbase TZ x331 = 0x2;
-Tbase TZ x332 = x97 * x331;
-Tbase TZ x333 = x91 * x332;
-Tbase TZ x334 = x92 * x96;
-Tbase TZ x335 = 0x2;
-Tbase TZ x336 = x95 * x335;
-Tbase TZ x337 = x93 * x336;
-Tbase TZ x338 = x94 * x94;
-Tbase TZ x339 = 0x2;
-Tbase TZ x340 = x93 * x339;
-Tbase TZ x341 = x95 * x340;
-Tbase TZ x342 = x96 * x92;
-Tbase TZ x343 = 0x2;
-Tbase TZ x344 = x91 * x343;
-Tbase TZ x345 = x97 * x344;
-Tbase TZ x346 = x98 * x90;
-Tbase TZ x347 = x345 + x346;
-Tbase TZ x348 = x342 + x347;
-Tbase TZ x349 = x341 + x348;
-Tbase TZ x350 = x338 + x349;
-Tbase TZ x351 = x337 + x350;
-Tbase TZ x352 = x334 + x351;
-Tbase TZ x353 = x333 + x352;
-Tbase TZ x354 = x330 + x353;
-Tbase TZ x355 = 0x2;
-Tbase TZ x356 = x89 * x355;
-Tbase TZ x357 = x89 * x356;
-Tbase TZ x358 = 0x13;
-Tbase TZ x359 = x358 * x357;
-Tbase TZ x360 = x354 + x359;
-Tbase TZ x361 = x329 + x360;
-Tbase TZ x362 = 0x1a;
-Tbase TZ x363 = x361 >>> x362;
-Tbase TZ x364 = x89 * x98;
-Tbase TZ x365 = x90 * x97;
-Tbase TZ x366 = x91 * x96;
-Tbase TZ x367 = x92 * x95;
-Tbase TZ x368 = x93 * x94;
-Tbase TZ x369 = x94 * x93;
-Tbase TZ x370 = x95 * x92;
-Tbase TZ x371 = x96 * x91;
-Tbase TZ x372 = x97 * x90;
-Tbase TZ x373 = x98 * x89;
-Tbase TZ x374 = x372 + x373;
-Tbase TZ x375 = x371 + x374;
-Tbase TZ x376 = x370 + x375;
-Tbase TZ x377 = x369 + x376;
-Tbase TZ x378 = x368 + x377;
-Tbase TZ x379 = x367 + x378;
-Tbase TZ x380 = x366 + x379;
-Tbase TZ x381 = x365 + x380;
-Tbase TZ x382 = x364 + x381;
-Tbase TZ x383 = x363 + x382;
-Tbase TZ x384 = 0x19;
-Tbase TZ x385 = x383 >>> x384;
-Tbase TZ x386 = 0x13;
-Tbase TZ x387 = x386 * x385;
-Tbase TZ x388 = 0x3ffffff;
-Tbase TZ x389 = x129 & x388;
-Tbase TZ x390 = x387 + x389;
-Tbase TZ x391 = 0x1a;
-Tbase TZ x392 = x390 >>> x391;
-Tbase TZ x393 = 0x1ffffff;
-Tbase TZ x394 = x153 & x393;
-Tbase TZ x395 = x392 + x394;
-Tbase TZ x396 = 0x1ffffff;
-Tbase TZ x397 = x383 & x396;
-Tbase TZ x398 = 0x3ffffff;
-Tbase TZ x399 = x361 & x398;
-Tbase TZ x400 = 0x1ffffff;
-Tbase TZ x401 = x327 & x400;
-Tbase TZ x402 = 0x3ffffff;
-Tbase TZ x403 = x303 & x402;
-Tbase TZ x404 = 0x1ffffff;
-Tbase TZ x405 = x269 & x404;
-Tbase TZ x406 = 0x3ffffff;
-Tbase TZ x407 = x245 & x406;
-Tbase TZ x408 = 0x1ffffff;
-Tbase TZ x409 = x211 & x408;
-Tbase TZ x410 = 0x19;
-Tbase TZ x411 = x395 >>> x410;
-Tbase TZ x412 = 0x3ffffff;
-Tbase TZ x413 = x187 & x412;
-Tbase TZ x414 = x411 + x413;
-Tbase TZ x415 = 0x1ffffff;
-Tbase TZ x416 = x395 & x415;
-Tbase TZ x417 = 0x3ffffff;
-Tbase TZ x418 = x390 & x417;
-Tbase TZ x419 = 0x3fffffe;
-Tbase TZ x420 = x419 + x49;
-Tbase TZ x421 = x420 - x59;
-Tbase TZ x422 = 0x7fffffe;
-Tbase TZ x423 = x422 + x50;
-Tbase TZ x424 = x423 - x60;
-Tbase TZ x425 = 0x3fffffe;
-Tbase TZ x426 = x425 + x51;
-Tbase TZ x427 = x426 - x61;
-Tbase TZ x428 = 0x7fffffe;
-Tbase TZ x429 = x428 + x52;
-Tbase TZ x430 = x429 - x62;
-Tbase TZ x431 = 0x3fffffe;
-Tbase TZ x432 = x431 + x53;
-Tbase TZ x433 = x432 - x63;
-Tbase TZ x434 = 0x7fffffe;
-Tbase TZ x435 = x434 + x54;
-Tbase TZ x436 = x435 - x64;
-Tbase TZ x437 = 0x3fffffe;
-Tbase TZ x438 = x437 + x55;
-Tbase TZ x439 = x438 - x65;
-Tbase TZ x440 = 0x7fffffe;
-Tbase TZ x441 = x440 + x56;
-Tbase TZ x442 = x441 - x66;
-Tbase TZ x443 = 0x3fffffe;
-Tbase TZ x444 = x443 + x57;
-Tbase TZ x445 = x444 - x67;
-Tbase TZ x446 = 0x7ffffda;
-Tbase TZ x447 = x446 + x58;
-Tbase TZ x448 = x447 - x68;
-Tbase TZ x449 = x448 * x448;
-Tbase TZ x450 = 0x2;
-Tbase TZ x451 = x445 * x450;
-Tbase TZ x452 = x421 * x451;
-Tbase TZ x453 = x424 * x442;
-Tbase TZ x454 = 0x2;
-Tbase TZ x455 = x439 * x454;
-Tbase TZ x456 = x427 * x455;
-Tbase TZ x457 = x430 * x436;
-Tbase TZ x458 = 0x2;
-Tbase TZ x459 = x433 * x458;
-Tbase TZ x460 = x433 * x459;
-Tbase TZ x461 = x436 * x430;
-Tbase TZ x462 = 0x2;
-Tbase TZ x463 = x427 * x462;
-Tbase TZ x464 = x439 * x463;
-Tbase TZ x465 = x442 * x424;
-Tbase TZ x466 = 0x2;
-Tbase TZ x467 = x421 * x466;
-Tbase TZ x468 = x445 * x467;
-Tbase TZ x469 = x465 + x468;
-Tbase TZ x470 = x464 + x469;
-Tbase TZ x471 = x461 + x470;
-Tbase TZ x472 = x460 + x471;
-Tbase TZ x473 = x457 + x472;
-Tbase TZ x474 = x456 + x473;
-Tbase TZ x475 = x453 + x474;
-Tbase TZ x476 = x452 + x475;
-Tbase TZ x477 = 0x13;
-Tbase TZ x478 = x477 * x476;
-Tbase TZ x479 = x449 + x478;
-Tbase TZ x480 = 0x1a;
-Tbase TZ x481 = x479 >>> x480;
-Tbase TZ x482 = x445 * x448;
-Tbase TZ x483 = x448 * x445;
-Tbase TZ x484 = x482 + x483;
-Tbase TZ x485 = x421 * x442;
-Tbase TZ x486 = x424 * x439;
-Tbase TZ x487 = x427 * x436;
-Tbase TZ x488 = x430 * x433;
-Tbase TZ x489 = x433 * x430;
-Tbase TZ x490 = x436 * x427;
-Tbase TZ x491 = x439 * x424;
-Tbase TZ x492 = x442 * x421;
-Tbase TZ x493 = x491 + x492;
-Tbase TZ x494 = x490 + x493;
-Tbase TZ x495 = x489 + x494;
-Tbase TZ x496 = x488 + x495;
-Tbase TZ x497 = x487 + x496;
-Tbase TZ x498 = x486 + x497;
-Tbase TZ x499 = x485 + x498;
-Tbase TZ x500 = 0x13;
-Tbase TZ x501 = x500 * x499;
-Tbase TZ x502 = x484 + x501;
-Tbase TZ x503 = x481 + x502;
-Tbase TZ x504 = 0x19;
-Tbase TZ x505 = x503 >>> x504;
-Tbase TZ x506 = x442 * x448;
-Tbase TZ x507 = 0x2;
-Tbase TZ x508 = x445 * x507;
-Tbase TZ x509 = x445 * x508;
-Tbase TZ x510 = x448 * x442;
-Tbase TZ x511 = x509 + x510;
-Tbase TZ x512 = x506 + x511;
-Tbase TZ x513 = 0x2;
-Tbase TZ x514 = x439 * x513;
-Tbase TZ x515 = x421 * x514;
-Tbase TZ x516 = x424 * x436;
-Tbase TZ x517 = 0x2;
-Tbase TZ x518 = x433 * x517;
-Tbase TZ x519 = x427 * x518;
-Tbase TZ x520 = x430 * x430;
-Tbase TZ x521 = 0x2;
-Tbase TZ x522 = x427 * x521;
-Tbase TZ x523 = x433 * x522;
-Tbase TZ x524 = x436 * x424;
-Tbase TZ x525 = 0x2;
-Tbase TZ x526 = x421 * x525;
-Tbase TZ x527 = x439 * x526;
-Tbase TZ x528 = x524 + x527;
-Tbase TZ x529 = x523 + x528;
-Tbase TZ x530 = x520 + x529;
-Tbase TZ x531 = x519 + x530;
-Tbase TZ x532 = x516 + x531;
-Tbase TZ x533 = x515 + x532;
-Tbase TZ x534 = 0x13;
-Tbase TZ x535 = x534 * x533;
-Tbase TZ x536 = x512 + x535;
-Tbase TZ x537 = x505 + x536;
-Tbase TZ x538 = 0x1a;
-Tbase TZ x539 = x537 >>> x538;
-Tbase TZ x540 = x439 * x448;
-Tbase TZ x541 = x442 * x445;
-Tbase TZ x542 = x445 * x442;
-Tbase TZ x543 = x448 * x439;
-Tbase TZ x544 = x542 + x543;
-Tbase TZ x545 = x541 + x544;
-Tbase TZ x546 = x540 + x545;
-Tbase TZ x547 = x421 * x436;
-Tbase TZ x548 = x424 * x433;
-Tbase TZ x549 = x427 * x430;
-Tbase TZ x550 = x430 * x427;
-Tbase TZ x551 = x433 * x424;
-Tbase TZ x552 = x436 * x421;
-Tbase TZ x553 = x551 + x552;
-Tbase TZ x554 = x550 + x553;
-Tbase TZ x555 = x549 + x554;
-Tbase TZ x556 = x548 + x555;
-Tbase TZ x557 = x547 + x556;
-Tbase TZ x558 = 0x13;
-Tbase TZ x559 = x558 * x557;
-Tbase TZ x560 = x546 + x559;
-Tbase TZ x561 = x539 + x560;
-Tbase TZ x562 = 0x19;
-Tbase TZ x563 = x561 >>> x562;
-Tbase TZ x564 = x436 * x448;
-Tbase TZ x565 = 0x2;
-Tbase TZ x566 = x445 * x565;
-Tbase TZ x567 = x439 * x566;
-Tbase TZ x568 = x442 * x442;
-Tbase TZ x569 = 0x2;
-Tbase TZ x570 = x439 * x569;
-Tbase TZ x571 = x445 * x570;
-Tbase TZ x572 = x448 * x436;
-Tbase TZ x573 = x571 + x572;
-Tbase TZ x574 = x568 + x573;
-Tbase TZ x575 = x567 + x574;
-Tbase TZ x576 = x564 + x575;
-Tbase TZ x577 = 0x2;
-Tbase TZ x578 = x433 * x577;
-Tbase TZ x579 = x421 * x578;
-Tbase TZ x580 = x424 * x430;
-Tbase TZ x581 = 0x2;
-Tbase TZ x582 = x427 * x581;
-Tbase TZ x583 = x427 * x582;
-Tbase TZ x584 = x430 * x424;
-Tbase TZ x585 = 0x2;
-Tbase TZ x586 = x421 * x585;
-Tbase TZ x587 = x433 * x586;
-Tbase TZ x588 = x584 + x587;
-Tbase TZ x589 = x583 + x588;
-Tbase TZ x590 = x580 + x589;
-Tbase TZ x591 = x579 + x590;
-Tbase TZ x592 = 0x13;
-Tbase TZ x593 = x592 * x591;
-Tbase TZ x594 = x576 + x593;
-Tbase TZ x595 = x563 + x594;
-Tbase TZ x596 = 0x1a;
-Tbase TZ x597 = x595 >>> x596;
-Tbase TZ x598 = x433 * x448;
-Tbase TZ x599 = x436 * x445;
-Tbase TZ x600 = x439 * x442;
-Tbase TZ x601 = x442 * x439;
-Tbase TZ x602 = x445 * x436;
-Tbase TZ x603 = x448 * x433;
-Tbase TZ x604 = x602 + x603;
-Tbase TZ x605 = x601 + x604;
-Tbase TZ x606 = x600 + x605;
-Tbase TZ x607 = x599 + x606;
-Tbase TZ x608 = x598 + x607;
-Tbase TZ x609 = x421 * x430;
-Tbase TZ x610 = x424 * x427;
-Tbase TZ x611 = x427 * x424;
-Tbase TZ x612 = x430 * x421;
-Tbase TZ x613 = x611 + x612;
-Tbase TZ x614 = x610 + x613;
-Tbase TZ x615 = x609 + x614;
-Tbase TZ x616 = 0x13;
-Tbase TZ x617 = x616 * x615;
-Tbase TZ x618 = x608 + x617;
-Tbase TZ x619 = x597 + x618;
-Tbase TZ x620 = 0x19;
-Tbase TZ x621 = x619 >>> x620;
-Tbase TZ x622 = x430 * x448;
-Tbase TZ x623 = 0x2;
-Tbase TZ x624 = x445 * x623;
-Tbase TZ x625 = x433 * x624;
-Tbase TZ x626 = x436 * x442;
-Tbase TZ x627 = 0x2;
-Tbase TZ x628 = x439 * x627;
-Tbase TZ x629 = x439 * x628;
-Tbase TZ x630 = x442 * x436;
-Tbase TZ x631 = 0x2;
-Tbase TZ x632 = x433 * x631;
-Tbase TZ x633 = x445 * x632;
-Tbase TZ x634 = x448 * x430;
-Tbase TZ x635 = x633 + x634;
-Tbase TZ x636 = x630 + x635;
-Tbase TZ x637 = x629 + x636;
-Tbase TZ x638 = x626 + x637;
-Tbase TZ x639 = x625 + x638;
-Tbase TZ x640 = x622 + x639;
-Tbase TZ x641 = 0x2;
-Tbase TZ x642 = x427 * x641;
-Tbase TZ x643 = x421 * x642;
-Tbase TZ x644 = x424 * x424;
-Tbase TZ x645 = 0x2;
-Tbase TZ x646 = x421 * x645;
-Tbase TZ x647 = x427 * x646;
-Tbase TZ x648 = x644 + x647;
-Tbase TZ x649 = x643 + x648;
-Tbase TZ x650 = 0x13;
-Tbase TZ x651 = x650 * x649;
-Tbase TZ x652 = x640 + x651;
-Tbase TZ x653 = x621 + x652;
-Tbase TZ x654 = 0x1a;
-Tbase TZ x655 = x653 >>> x654;
-Tbase TZ x656 = x427 * x448;
-Tbase TZ x657 = x430 * x445;
-Tbase TZ x658 = x433 * x442;
-Tbase TZ x659 = x436 * x439;
-Tbase TZ x660 = x439 * x436;
-Tbase TZ x661 = x442 * x433;
-Tbase TZ x662 = x445 * x430;
-Tbase TZ x663 = x448 * x427;
-Tbase TZ x664 = x662 + x663;
-Tbase TZ x665 = x661 + x664;
-Tbase TZ x666 = x660 + x665;
-Tbase TZ x667 = x659 + x666;
-Tbase TZ x668 = x658 + x667;
-Tbase TZ x669 = x657 + x668;
-Tbase TZ x670 = x656 + x669;
-Tbase TZ x671 = x421 * x424;
-Tbase TZ x672 = x424 * x421;
-Tbase TZ x673 = x671 + x672;
-Tbase TZ x674 = 0x13;
-Tbase TZ x675 = x674 * x673;
-Tbase TZ x676 = x670 + x675;
-Tbase TZ x677 = x655 + x676;
-Tbase TZ x678 = 0x19;
-Tbase TZ x679 = x677 >>> x678;
-Tbase TZ x680 = x424 * x448;
-Tbase TZ x681 = 0x2;
-Tbase TZ x682 = x445 * x681;
-Tbase TZ x683 = x427 * x682;
-Tbase TZ x684 = x430 * x442;
-Tbase TZ x685 = 0x2;
-Tbase TZ x686 = x439 * x685;
-Tbase TZ x687 = x433 * x686;
-Tbase TZ x688 = x436 * x436;
-Tbase TZ x689 = 0x2;
-Tbase TZ x690 = x433 * x689;
-Tbase TZ x691 = x439 * x690;
-Tbase TZ x692 = x442 * x430;
-Tbase TZ x693 = 0x2;
-Tbase TZ x694 = x427 * x693;
-Tbase TZ x695 = x445 * x694;
-Tbase TZ x696 = x448 * x424;
-Tbase TZ x697 = x695 + x696;
-Tbase TZ x698 = x692 + x697;
-Tbase TZ x699 = x691 + x698;
-Tbase TZ x700 = x688 + x699;
-Tbase TZ x701 = x687 + x700;
-Tbase TZ x702 = x684 + x701;
-Tbase TZ x703 = x683 + x702;
-Tbase TZ x704 = x680 + x703;
-Tbase TZ x705 = 0x2;
-Tbase TZ x706 = x421 * x705;
-Tbase TZ x707 = x421 * x706;
-Tbase TZ x708 = 0x13;
-Tbase TZ x709 = x708 * x707;
-Tbase TZ x710 = x704 + x709;
-Tbase TZ x711 = x679 + x710;
-Tbase TZ x712 = 0x1a;
-Tbase TZ x713 = x711 >>> x712;
-Tbase TZ x714 = x421 * x448;
-Tbase TZ x715 = x424 * x445;
-Tbase TZ x716 = x427 * x442;
-Tbase TZ x717 = x430 * x439;
-Tbase TZ x718 = x433 * x436;
-Tbase TZ x719 = x436 * x433;
-Tbase TZ x720 = x439 * x430;
-Tbase TZ x721 = x442 * x427;
-Tbase TZ x722 = x445 * x424;
-Tbase TZ x723 = x448 * x421;
-Tbase TZ x724 = x722 + x723;
-Tbase TZ x725 = x721 + x724;
-Tbase TZ x726 = x720 + x725;
-Tbase TZ x727 = x719 + x726;
-Tbase TZ x728 = x718 + x727;
-Tbase TZ x729 = x717 + x728;
-Tbase TZ x730 = x716 + x729;
-Tbase TZ x731 = x715 + x730;
-Tbase TZ x732 = x714 + x731;
-Tbase TZ x733 = x713 + x732;
-Tbase TZ x734 = 0x19;
-Tbase TZ x735 = x733 >>> x734;
-Tbase TZ x736 = 0x13;
-Tbase TZ x737 = x736 * x735;
-Tbase TZ x738 = 0x3ffffff;
-Tbase TZ x739 = x479 & x738;
-Tbase TZ x740 = x737 + x739;
-Tbase TZ x741 = 0x1a;
-Tbase TZ x742 = x740 >>> x741;
-Tbase TZ x743 = 0x1ffffff;
-Tbase TZ x744 = x503 & x743;
-Tbase TZ x745 = x742 + x744;
-Tbase TZ x746 = 0x1ffffff;
-Tbase TZ x747 = x733 & x746;
-Tbase TZ x748 = 0x3ffffff;
-Tbase TZ x749 = x711 & x748;
-Tbase TZ x750 = 0x1ffffff;
-Tbase TZ x751 = x677 & x750;
-Tbase TZ x752 = 0x3ffffff;
-Tbase TZ x753 = x653 & x752;
-Tbase TZ x754 = 0x1ffffff;
-Tbase TZ x755 = x619 & x754;
-Tbase TZ x756 = 0x3ffffff;
-Tbase TZ x757 = x595 & x756;
-Tbase TZ x758 = 0x1ffffff;
-Tbase TZ x759 = x561 & x758;
-Tbase TZ x760 = 0x19;
-Tbase TZ x761 = x745 >>> x760;
-Tbase TZ x762 = 0x3ffffff;
-Tbase TZ x763 = x537 & x762;
-Tbase TZ x764 = x761 + x763;
-Tbase TZ x765 = 0x1ffffff;
-Tbase TZ x766 = x745 & x765;
-Tbase TZ x767 = 0x3ffffff;
-Tbase TZ x768 = x740 & x767;
-Tbase TZ x769 = 0x3fffffe;
-Tbase TZ x770 = x769 + x397;
-Tbase TZ x771 = x770 - x747;
-Tbase TZ x772 = 0x7fffffe;
-Tbase TZ x773 = x772 + x399;
-Tbase TZ x774 = x773 - x749;
-Tbase TZ x775 = 0x3fffffe;
-Tbase TZ x776 = x775 + x401;
-Tbase TZ x777 = x776 - x751;
-Tbase TZ x778 = 0x7fffffe;
-Tbase TZ x779 = x778 + x403;
-Tbase TZ x780 = x779 - x753;
-Tbase TZ x781 = 0x3fffffe;
-Tbase TZ x782 = x781 + x405;
-Tbase TZ x783 = x782 - x755;
-Tbase TZ x784 = 0x7fffffe;
-Tbase TZ x785 = x784 + x407;
-Tbase TZ x786 = x785 - x757;
-Tbase TZ x787 = 0x3fffffe;
-Tbase TZ x788 = x787 + x409;
-Tbase TZ x789 = x788 - x759;
-Tbase TZ x790 = 0x7fffffe;
-Tbase TZ x791 = x790 + x414;
-Tbase TZ x792 = x791 - x764;
-Tbase TZ x793 = 0x3fffffe;
-Tbase TZ x794 = x793 + x416;
-Tbase TZ x795 = x794 - x766;
-Tbase TZ x796 = 0x7ffffda;
-Tbase TZ x797 = x796 + x418;
-Tbase TZ x798 = x797 - x768;
-Tbase TZ x799 = x69 + x79;
-Tbase TZ x800 = x70 + x80;
-Tbase TZ x801 = x71 + x81;
-Tbase TZ x802 = x72 + x82;
-Tbase TZ x803 = x73 + x83;
-Tbase TZ x804 = x74 + x84;
-Tbase TZ x805 = x75 + x85;
-Tbase TZ x806 = x76 + x86;
-Tbase TZ x807 = x77 + x87;
-Tbase TZ x808 = x78 + x88;
-Tbase TZ x809 = 0x3fffffe;
-Tbase TZ x810 = x809 + x69;
-Tbase TZ x811 = x810 - x79;
-Tbase TZ x812 = 0x7fffffe;
-Tbase TZ x813 = x812 + x70;
-Tbase TZ x814 = x813 - x80;
-Tbase TZ x815 = 0x3fffffe;
-Tbase TZ x816 = x815 + x71;
-Tbase TZ x817 = x816 - x81;
-Tbase TZ x818 = 0x7fffffe;
-Tbase TZ x819 = x818 + x72;
-Tbase TZ x820 = x819 - x82;
-Tbase TZ x821 = 0x3fffffe;
-Tbase TZ x822 = x821 + x73;
-Tbase TZ x823 = x822 - x83;
-Tbase TZ x824 = 0x7fffffe;
-Tbase TZ x825 = x824 + x74;
-Tbase TZ x826 = x825 - x84;
-Tbase TZ x827 = 0x3fffffe;
-Tbase TZ x828 = x827 + x75;
-Tbase TZ x829 = x828 - x85;
-Tbase TZ x830 = 0x7fffffe;
-Tbase TZ x831 = x830 + x76;
-Tbase TZ x832 = x831 - x86;
-Tbase TZ x833 = 0x3fffffe;
-Tbase TZ x834 = x833 + x77;
-Tbase TZ x835 = x834 - x87;
-Tbase TZ x836 = 0x7ffffda;
-Tbase TZ x837 = x836 + x78;
-Tbase TZ x838 = x837 - x88;
-Tbase TZ x839 = x838 * x98;
-Tbase TZ x840 = 0x2;
-Tbase TZ x841 = x97 * x840;
-Tbase TZ x842 = x811 * x841;
-Tbase TZ x843 = x814 * x96;
-Tbase TZ x844 = 0x2;
-Tbase TZ x845 = x95 * x844;
-Tbase TZ x846 = x817 * x845;
-Tbase TZ x847 = x820 * x94;
-Tbase TZ x848 = 0x2;
-Tbase TZ x849 = x93 * x848;
-Tbase TZ x850 = x823 * x849;
-Tbase TZ x851 = x826 * x92;
-Tbase TZ x852 = 0x2;
-Tbase TZ x853 = x91 * x852;
-Tbase TZ x854 = x829 * x853;
-Tbase TZ x855 = x832 * x90;
-Tbase TZ x856 = 0x2;
-Tbase TZ x857 = x89 * x856;
-Tbase TZ x858 = x835 * x857;
-Tbase TZ x859 = x855 + x858;
-Tbase TZ x860 = x854 + x859;
-Tbase TZ x861 = x851 + x860;
-Tbase TZ x862 = x850 + x861;
-Tbase TZ x863 = x847 + x862;
-Tbase TZ x864 = x846 + x863;
-Tbase TZ x865 = x843 + x864;
-Tbase TZ x866 = x842 + x865;
-Tbase TZ x867 = 0x13;
-Tbase TZ x868 = x867 * x866;
-Tbase TZ x869 = x839 + x868;
-Tbase TZ x870 = 0x1a;
-Tbase TZ x871 = x869 >>> x870;
-Tbase TZ x872 = x835 * x98;
-Tbase TZ x873 = x838 * x97;
-Tbase TZ x874 = x872 + x873;
-Tbase TZ x875 = x811 * x96;
-Tbase TZ x876 = x814 * x95;
-Tbase TZ x877 = x817 * x94;
-Tbase TZ x878 = x820 * x93;
-Tbase TZ x879 = x823 * x92;
-Tbase TZ x880 = x826 * x91;
-Tbase TZ x881 = x829 * x90;
-Tbase TZ x882 = x832 * x89;
-Tbase TZ x883 = x881 + x882;
-Tbase TZ x884 = x880 + x883;
-Tbase TZ x885 = x879 + x884;
-Tbase TZ x886 = x878 + x885;
-Tbase TZ x887 = x877 + x886;
-Tbase TZ x888 = x876 + x887;
-Tbase TZ x889 = x875 + x888;
-Tbase TZ x890 = 0x13;
-Tbase TZ x891 = x890 * x889;
-Tbase TZ x892 = x874 + x891;
-Tbase TZ x893 = x871 + x892;
-Tbase TZ x894 = 0x19;
-Tbase TZ x895 = x893 >>> x894;
-Tbase TZ x896 = x832 * x98;
-Tbase TZ x897 = 0x2;
-Tbase TZ x898 = x97 * x897;
-Tbase TZ x899 = x835 * x898;
-Tbase TZ x900 = x838 * x96;
-Tbase TZ x901 = x899 + x900;
-Tbase TZ x902 = x896 + x901;
-Tbase TZ x903 = 0x2;
-Tbase TZ x904 = x95 * x903;
-Tbase TZ x905 = x811 * x904;
-Tbase TZ x906 = x814 * x94;
-Tbase TZ x907 = 0x2;
-Tbase TZ x908 = x93 * x907;
-Tbase TZ x909 = x817 * x908;
-Tbase TZ x910 = x820 * x92;
-Tbase TZ x911 = 0x2;
-Tbase TZ x912 = x91 * x911;
-Tbase TZ x913 = x823 * x912;
-Tbase TZ x914 = x826 * x90;
-Tbase TZ x915 = 0x2;
-Tbase TZ x916 = x89 * x915;
-Tbase TZ x917 = x829 * x916;
-Tbase TZ x918 = x914 + x917;
-Tbase TZ x919 = x913 + x918;
-Tbase TZ x920 = x910 + x919;
-Tbase TZ x921 = x909 + x920;
-Tbase TZ x922 = x906 + x921;
-Tbase TZ x923 = x905 + x922;
-Tbase TZ x924 = 0x13;
-Tbase TZ x925 = x924 * x923;
-Tbase TZ x926 = x902 + x925;
-Tbase TZ x927 = x895 + x926;
-Tbase TZ x928 = 0x1a;
-Tbase TZ x929 = x927 >>> x928;
-Tbase TZ x930 = x829 * x98;
-Tbase TZ x931 = x832 * x97;
-Tbase TZ x932 = x835 * x96;
-Tbase TZ x933 = x838 * x95;
-Tbase TZ x934 = x932 + x933;
-Tbase TZ x935 = x931 + x934;
-Tbase TZ x936 = x930 + x935;
-Tbase TZ x937 = x811 * x94;
-Tbase TZ x938 = x814 * x93;
-Tbase TZ x939 = x817 * x92;
-Tbase TZ x940 = x820 * x91;
-Tbase TZ x941 = x823 * x90;
-Tbase TZ x942 = x826 * x89;
-Tbase TZ x943 = x941 + x942;
-Tbase TZ x944 = x940 + x943;
-Tbase TZ x945 = x939 + x944;
-Tbase TZ x946 = x938 + x945;
-Tbase TZ x947 = x937 + x946;
-Tbase TZ x948 = 0x13;
-Tbase TZ x949 = x948 * x947;
-Tbase TZ x950 = x936 + x949;
-Tbase TZ x951 = x929 + x950;
-Tbase TZ x952 = 0x19;
-Tbase TZ x953 = x951 >>> x952;
-Tbase TZ x954 = x826 * x98;
-Tbase TZ x955 = 0x2;
-Tbase TZ x956 = x97 * x955;
-Tbase TZ x957 = x829 * x956;
-Tbase TZ x958 = x832 * x96;
-Tbase TZ x959 = 0x2;
-Tbase TZ x960 = x95 * x959;
-Tbase TZ x961 = x835 * x960;
-Tbase TZ x962 = x838 * x94;
-Tbase TZ x963 = x961 + x962;
-Tbase TZ x964 = x958 + x963;
-Tbase TZ x965 = x957 + x964;
-Tbase TZ x966 = x954 + x965;
-Tbase TZ x967 = 0x2;
-Tbase TZ x968 = x93 * x967;
-Tbase TZ x969 = x811 * x968;
-Tbase TZ x970 = x814 * x92;
-Tbase TZ x971 = 0x2;
-Tbase TZ x972 = x91 * x971;
-Tbase TZ x973 = x817 * x972;
-Tbase TZ x974 = x820 * x90;
-Tbase TZ x975 = 0x2;
-Tbase TZ x976 = x89 * x975;
-Tbase TZ x977 = x823 * x976;
-Tbase TZ x978 = x974 + x977;
-Tbase TZ x979 = x973 + x978;
-Tbase TZ x980 = x970 + x979;
-Tbase TZ x981 = x969 + x980;
-Tbase TZ x982 = 0x13;
-Tbase TZ x983 = x982 * x981;
-Tbase TZ x984 = x966 + x983;
-Tbase TZ x985 = x953 + x984;
-Tbase TZ x986 = 0x1a;
-Tbase TZ x987 = x985 >>> x986;
-Tbase TZ x988 = x823 * x98;
-Tbase TZ x989 = x826 * x97;
-Tbase TZ x990 = x829 * x96;
-Tbase TZ x991 = x832 * x95;
-Tbase TZ x992 = x835 * x94;
-Tbase TZ x993 = x838 * x93;
-Tbase TZ x994 = x992 + x993;
-Tbase TZ x995 = x991 + x994;
-Tbase TZ x996 = x990 + x995;
-Tbase TZ x997 = x989 + x996;
-Tbase TZ x998 = x988 + x997;
-Tbase TZ x999 = x811 * x92;
-Tbase TZ x1000 = x814 * x91;
-Tbase TZ x1001 = x817 * x90;
-Tbase TZ x1002 = x820 * x89;
-Tbase TZ x1003 = x1001 + x1002;
-Tbase TZ x1004 = x1000 + x1003;
-Tbase TZ x1005 = x999 + x1004;
-Tbase TZ x1006 = 0x13;
-Tbase TZ x1007 = x1006 * x1005;
-Tbase TZ x1008 = x998 + x1007;
-Tbase TZ x1009 = x987 + x1008;
-Tbase TZ x1010 = 0x19;
-Tbase TZ x1011 = x1009 >>> x1010;
-Tbase TZ x1012 = x820 * x98;
-Tbase TZ x1013 = 0x2;
-Tbase TZ x1014 = x97 * x1013;
-Tbase TZ x1015 = x823 * x1014;
-Tbase TZ x1016 = x826 * x96;
-Tbase TZ x1017 = 0x2;
-Tbase TZ x1018 = x95 * x1017;
-Tbase TZ x1019 = x829 * x1018;
-Tbase TZ x1020 = x832 * x94;
-Tbase TZ x1021 = 0x2;
-Tbase TZ x1022 = x93 * x1021;
-Tbase TZ x1023 = x835 * x1022;
-Tbase TZ x1024 = x838 * x92;
-Tbase TZ x1025 = x1023 + x1024;
-Tbase TZ x1026 = x1020 + x1025;
-Tbase TZ x1027 = x1019 + x1026;
-Tbase TZ x1028 = x1016 + x1027;
-Tbase TZ x1029 = x1015 + x1028;
-Tbase TZ x1030 = x1012 + x1029;
-Tbase TZ x1031 = 0x2;
-Tbase TZ x1032 = x91 * x1031;
-Tbase TZ x1033 = x811 * x1032;
-Tbase TZ x1034 = x814 * x90;
-Tbase TZ x1035 = 0x2;
-Tbase TZ x1036 = x89 * x1035;
-Tbase TZ x1037 = x817 * x1036;
-Tbase TZ x1038 = x1034 + x1037;
-Tbase TZ x1039 = x1033 + x1038;
-Tbase TZ x1040 = 0x13;
-Tbase TZ x1041 = x1040 * x1039;
-Tbase TZ x1042 = x1030 + x1041;
-Tbase TZ x1043 = x1011 + x1042;
-Tbase TZ x1044 = 0x1a;
-Tbase TZ x1045 = x1043 >>> x1044;
-Tbase TZ x1046 = x817 * x98;
-Tbase TZ x1047 = x820 * x97;
-Tbase TZ x1048 = x823 * x96;
-Tbase TZ x1049 = x826 * x95;
-Tbase TZ x1050 = x829 * x94;
-Tbase TZ x1051 = x832 * x93;
-Tbase TZ x1052 = x835 * x92;
-Tbase TZ x1053 = x838 * x91;
-Tbase TZ x1054 = x1052 + x1053;
-Tbase TZ x1055 = x1051 + x1054;
-Tbase TZ x1056 = x1050 + x1055;
-Tbase TZ x1057 = x1049 + x1056;
-Tbase TZ x1058 = x1048 + x1057;
-Tbase TZ x1059 = x1047 + x1058;
-Tbase TZ x1060 = x1046 + x1059;
-Tbase TZ x1061 = x811 * x90;
-Tbase TZ x1062 = x814 * x89;
-Tbase TZ x1063 = x1061 + x1062;
-Tbase TZ x1064 = 0x13;
-Tbase TZ x1065 = x1064 * x1063;
-Tbase TZ x1066 = x1060 + x1065;
-Tbase TZ x1067 = x1045 + x1066;
-Tbase TZ x1068 = 0x19;
-Tbase TZ x1069 = x1067 >>> x1068;
-Tbase TZ x1070 = x814 * x98;
-Tbase TZ x1071 = 0x2;
-Tbase TZ x1072 = x97 * x1071;
-Tbase TZ x1073 = x817 * x1072;
-Tbase TZ x1074 = x820 * x96;
-Tbase TZ x1075 = 0x2;
-Tbase TZ x1076 = x95 * x1075;
-Tbase TZ x1077 = x823 * x1076;
-Tbase TZ x1078 = x826 * x94;
-Tbase TZ x1079 = 0x2;
-Tbase TZ x1080 = x93 * x1079;
-Tbase TZ x1081 = x829 * x1080;
-Tbase TZ x1082 = x832 * x92;
-Tbase TZ x1083 = 0x2;
-Tbase TZ x1084 = x91 * x1083;
-Tbase TZ x1085 = x835 * x1084;
-Tbase TZ x1086 = x838 * x90;
-Tbase TZ x1087 = x1085 + x1086;
-Tbase TZ x1088 = x1082 + x1087;
-Tbase TZ x1089 = x1081 + x1088;
-Tbase TZ x1090 = x1078 + x1089;
-Tbase TZ x1091 = x1077 + x1090;
-Tbase TZ x1092 = x1074 + x1091;
-Tbase TZ x1093 = x1073 + x1092;
-Tbase TZ x1094 = x1070 + x1093;
-Tbase TZ x1095 = 0x2;
-Tbase TZ x1096 = x89 * x1095;
-Tbase TZ x1097 = x811 * x1096;
-Tbase TZ x1098 = 0x13;
-Tbase TZ x1099 = x1098 * x1097;
-Tbase TZ x1100 = x1094 + x1099;
-Tbase TZ x1101 = x1069 + x1100;
-Tbase TZ x1102 = 0x1a;
-Tbase TZ x1103 = x1101 >>> x1102;
-Tbase TZ x1104 = x811 * x98;
-Tbase TZ x1105 = x814 * x97;
-Tbase TZ x1106 = x817 * x96;
-Tbase TZ x1107 = x820 * x95;
-Tbase TZ x1108 = x823 * x94;
-Tbase TZ x1109 = x826 * x93;
-Tbase TZ x1110 = x829 * x92;
-Tbase TZ x1111 = x832 * x91;
-Tbase TZ x1112 = x835 * x90;
-Tbase TZ x1113 = x838 * x89;
-Tbase TZ x1114 = x1112 + x1113;
-Tbase TZ x1115 = x1111 + x1114;
-Tbase TZ x1116 = x1110 + x1115;
-Tbase TZ x1117 = x1109 + x1116;
-Tbase TZ x1118 = x1108 + x1117;
-Tbase TZ x1119 = x1107 + x1118;
-Tbase TZ x1120 = x1106 + x1119;
-Tbase TZ x1121 = x1105 + x1120;
-Tbase TZ x1122 = x1104 + x1121;
-Tbase TZ x1123 = x1103 + x1122;
-Tbase TZ x1124 = 0x19;
-Tbase TZ x1125 = x1123 >>> x1124;
-Tbase TZ x1126 = 0x13;
-Tbase TZ x1127 = x1126 * x1125;
-Tbase TZ x1128 = 0x3ffffff;
-Tbase TZ x1129 = x869 & x1128;
-Tbase TZ x1130 = x1127 + x1129;
-Tbase TZ x1131 = 0x1a;
-Tbase TZ x1132 = x1130 >>> x1131;
-Tbase TZ x1133 = 0x1ffffff;
-Tbase TZ x1134 = x893 & x1133;
-Tbase TZ x1135 = x1132 + x1134;
-Tbase TZ x1136 = 0x1ffffff;
-Tbase TZ x1137 = x1123 & x1136;
-Tbase TZ x1138 = 0x3ffffff;
-Tbase TZ x1139 = x1101 & x1138;
-Tbase TZ x1140 = 0x1ffffff;
-Tbase TZ x1141 = x1067 & x1140;
-Tbase TZ x1142 = 0x3ffffff;
-Tbase TZ x1143 = x1043 & x1142;
-Tbase TZ x1144 = 0x1ffffff;
-Tbase TZ x1145 = x1009 & x1144;
-Tbase TZ x1146 = 0x3ffffff;
-Tbase TZ x1147 = x985 & x1146;
-Tbase TZ x1148 = 0x1ffffff;
-Tbase TZ x1149 = x951 & x1148;
-Tbase TZ x1150 = 0x19;
-Tbase TZ x1151 = x1135 >>> x1150;
-Tbase TZ x1152 = 0x3ffffff;
-Tbase TZ x1153 = x927 & x1152;
-Tbase TZ x1154 = x1151 + x1153;
-Tbase TZ x1155 = 0x1ffffff;
-Tbase TZ x1156 = x1135 & x1155;
-Tbase TZ x1157 = 0x3ffffff;
-Tbase TZ x1158 = x1130 & x1157;
-Tbase TZ x1159 = x808 * x448;
-Tbase TZ x1160 = 0x2;
-Tbase TZ x1161 = x445 * x1160;
-Tbase TZ x1162 = x799 * x1161;
-Tbase TZ x1163 = x800 * x442;
-Tbase TZ x1164 = 0x2;
-Tbase TZ x1165 = x439 * x1164;
-Tbase TZ x1166 = x801 * x1165;
-Tbase TZ x1167 = x802 * x436;
-Tbase TZ x1168 = 0x2;
-Tbase TZ x1169 = x433 * x1168;
-Tbase TZ x1170 = x803 * x1169;
-Tbase TZ x1171 = x804 * x430;
-Tbase TZ x1172 = 0x2;
-Tbase TZ x1173 = x427 * x1172;
-Tbase TZ x1174 = x805 * x1173;
-Tbase TZ x1175 = x806 * x424;
-Tbase TZ x1176 = 0x2;
-Tbase TZ x1177 = x421 * x1176;
-Tbase TZ x1178 = x807 * x1177;
-Tbase TZ x1179 = x1175 + x1178;
-Tbase TZ x1180 = x1174 + x1179;
-Tbase TZ x1181 = x1171 + x1180;
-Tbase TZ x1182 = x1170 + x1181;
-Tbase TZ x1183 = x1167 + x1182;
-Tbase TZ x1184 = x1166 + x1183;
-Tbase TZ x1185 = x1163 + x1184;
-Tbase TZ x1186 = x1162 + x1185;
-Tbase TZ x1187 = 0x13;
-Tbase TZ x1188 = x1187 * x1186;
-Tbase TZ x1189 = x1159 + x1188;
-Tbase TZ x1190 = 0x1a;
-Tbase TZ x1191 = x1189 >>> x1190;
-Tbase TZ x1192 = x807 * x448;
-Tbase TZ x1193 = x808 * x445;
-Tbase TZ x1194 = x1192 + x1193;
-Tbase TZ x1195 = x799 * x442;
-Tbase TZ x1196 = x800 * x439;
-Tbase TZ x1197 = x801 * x436;
-Tbase TZ x1198 = x802 * x433;
-Tbase TZ x1199 = x803 * x430;
-Tbase TZ x1200 = x804 * x427;
-Tbase TZ x1201 = x805 * x424;
-Tbase TZ x1202 = x806 * x421;
-Tbase TZ x1203 = x1201 + x1202;
-Tbase TZ x1204 = x1200 + x1203;
-Tbase TZ x1205 = x1199 + x1204;
-Tbase TZ x1206 = x1198 + x1205;
-Tbase TZ x1207 = x1197 + x1206;
-Tbase TZ x1208 = x1196 + x1207;
-Tbase TZ x1209 = x1195 + x1208;
-Tbase TZ x1210 = 0x13;
-Tbase TZ x1211 = x1210 * x1209;
-Tbase TZ x1212 = x1194 + x1211;
-Tbase TZ x1213 = x1191 + x1212;
-Tbase TZ x1214 = 0x19;
-Tbase TZ x1215 = x1213 >>> x1214;
-Tbase TZ x1216 = x806 * x448;
-Tbase TZ x1217 = 0x2;
-Tbase TZ x1218 = x445 * x1217;
-Tbase TZ x1219 = x807 * x1218;
-Tbase TZ x1220 = x808 * x442;
-Tbase TZ x1221 = x1219 + x1220;
-Tbase TZ x1222 = x1216 + x1221;
-Tbase TZ x1223 = 0x2;
-Tbase TZ x1224 = x439 * x1223;
-Tbase TZ x1225 = x799 * x1224;
-Tbase TZ x1226 = x800 * x436;
-Tbase TZ x1227 = 0x2;
-Tbase TZ x1228 = x433 * x1227;
-Tbase TZ x1229 = x801 * x1228;
-Tbase TZ x1230 = x802 * x430;
-Tbase TZ x1231 = 0x2;
-Tbase TZ x1232 = x427 * x1231;
-Tbase TZ x1233 = x803 * x1232;
-Tbase TZ x1234 = x804 * x424;
-Tbase TZ x1235 = 0x2;
-Tbase TZ x1236 = x421 * x1235;
-Tbase TZ x1237 = x805 * x1236;
-Tbase TZ x1238 = x1234 + x1237;
-Tbase TZ x1239 = x1233 + x1238;
-Tbase TZ x1240 = x1230 + x1239;
-Tbase TZ x1241 = x1229 + x1240;
-Tbase TZ x1242 = x1226 + x1241;
-Tbase TZ x1243 = x1225 + x1242;
-Tbase TZ x1244 = 0x13;
-Tbase TZ x1245 = x1244 * x1243;
-Tbase TZ x1246 = x1222 + x1245;
-Tbase TZ x1247 = x1215 + x1246;
-Tbase TZ x1248 = 0x1a;
-Tbase TZ x1249 = x1247 >>> x1248;
-Tbase TZ x1250 = x805 * x448;
-Tbase TZ x1251 = x806 * x445;
-Tbase TZ x1252 = x807 * x442;
-Tbase TZ x1253 = x808 * x439;
-Tbase TZ x1254 = x1252 + x1253;
-Tbase TZ x1255 = x1251 + x1254;
-Tbase TZ x1256 = x1250 + x1255;
-Tbase TZ x1257 = x799 * x436;
-Tbase TZ x1258 = x800 * x433;
-Tbase TZ x1259 = x801 * x430;
-Tbase TZ x1260 = x802 * x427;
-Tbase TZ x1261 = x803 * x424;
-Tbase TZ x1262 = x804 * x421;
-Tbase TZ x1263 = x1261 + x1262;
-Tbase TZ x1264 = x1260 + x1263;
-Tbase TZ x1265 = x1259 + x1264;
-Tbase TZ x1266 = x1258 + x1265;
-Tbase TZ x1267 = x1257 + x1266;
-Tbase TZ x1268 = 0x13;
-Tbase TZ x1269 = x1268 * x1267;
-Tbase TZ x1270 = x1256 + x1269;
-Tbase TZ x1271 = x1249 + x1270;
-Tbase TZ x1272 = 0x19;
-Tbase TZ x1273 = x1271 >>> x1272;
-Tbase TZ x1274 = x804 * x448;
-Tbase TZ x1275 = 0x2;
-Tbase TZ x1276 = x445 * x1275;
-Tbase TZ x1277 = x805 * x1276;
-Tbase TZ x1278 = x806 * x442;
-Tbase TZ x1279 = 0x2;
-Tbase TZ x1280 = x439 * x1279;
-Tbase TZ x1281 = x807 * x1280;
-Tbase TZ x1282 = x808 * x436;
-Tbase TZ x1283 = x1281 + x1282;
-Tbase TZ x1284 = x1278 + x1283;
-Tbase TZ x1285 = x1277 + x1284;
-Tbase TZ x1286 = x1274 + x1285;
-Tbase TZ x1287 = 0x2;
-Tbase TZ x1288 = x433 * x1287;
-Tbase TZ x1289 = x799 * x1288;
-Tbase TZ x1290 = x800 * x430;
-Tbase TZ x1291 = 0x2;
-Tbase TZ x1292 = x427 * x1291;
-Tbase TZ x1293 = x801 * x1292;
-Tbase TZ x1294 = x802 * x424;
-Tbase TZ x1295 = 0x2;
-Tbase TZ x1296 = x421 * x1295;
-Tbase TZ x1297 = x803 * x1296;
-Tbase TZ x1298 = x1294 + x1297;
-Tbase TZ x1299 = x1293 + x1298;
-Tbase TZ x1300 = x1290 + x1299;
-Tbase TZ x1301 = x1289 + x1300;
-Tbase TZ x1302 = 0x13;
-Tbase TZ x1303 = x1302 * x1301;
-Tbase TZ x1304 = x1286 + x1303;
-Tbase TZ x1305 = x1273 + x1304;
-Tbase TZ x1306 = 0x1a;
-Tbase TZ x1307 = x1305 >>> x1306;
-Tbase TZ x1308 = x803 * x448;
-Tbase TZ x1309 = x804 * x445;
-Tbase TZ x1310 = x805 * x442;
-Tbase TZ x1311 = x806 * x439;
-Tbase TZ x1312 = x807 * x436;
-Tbase TZ x1313 = x808 * x433;
-Tbase TZ x1314 = x1312 + x1313;
-Tbase TZ x1315 = x1311 + x1314;
-Tbase TZ x1316 = x1310 + x1315;
-Tbase TZ x1317 = x1309 + x1316;
-Tbase TZ x1318 = x1308 + x1317;
-Tbase TZ x1319 = x799 * x430;
-Tbase TZ x1320 = x800 * x427;
-Tbase TZ x1321 = x801 * x424;
-Tbase TZ x1322 = x802 * x421;
-Tbase TZ x1323 = x1321 + x1322;
-Tbase TZ x1324 = x1320 + x1323;
-Tbase TZ x1325 = x1319 + x1324;
-Tbase TZ x1326 = 0x13;
-Tbase TZ x1327 = x1326 * x1325;
-Tbase TZ x1328 = x1318 + x1327;
-Tbase TZ x1329 = x1307 + x1328;
-Tbase TZ x1330 = 0x19;
-Tbase TZ x1331 = x1329 >>> x1330;
-Tbase TZ x1332 = x802 * x448;
-Tbase TZ x1333 = 0x2;
-Tbase TZ x1334 = x445 * x1333;
-Tbase TZ x1335 = x803 * x1334;
-Tbase TZ x1336 = x804 * x442;
-Tbase TZ x1337 = 0x2;
-Tbase TZ x1338 = x439 * x1337;
-Tbase TZ x1339 = x805 * x1338;
-Tbase TZ x1340 = x806 * x436;
-Tbase TZ x1341 = 0x2;
-Tbase TZ x1342 = x433 * x1341;
-Tbase TZ x1343 = x807 * x1342;
-Tbase TZ x1344 = x808 * x430;
-Tbase TZ x1345 = x1343 + x1344;
-Tbase TZ x1346 = x1340 + x1345;
-Tbase TZ x1347 = x1339 + x1346;
-Tbase TZ x1348 = x1336 + x1347;
-Tbase TZ x1349 = x1335 + x1348;
-Tbase TZ x1350 = x1332 + x1349;
-Tbase TZ x1351 = 0x2;
-Tbase TZ x1352 = x427 * x1351;
-Tbase TZ x1353 = x799 * x1352;
-Tbase TZ x1354 = x800 * x424;
-Tbase TZ x1355 = 0x2;
-Tbase TZ x1356 = x421 * x1355;
-Tbase TZ x1357 = x801 * x1356;
-Tbase TZ x1358 = x1354 + x1357;
-Tbase TZ x1359 = x1353 + x1358;
-Tbase TZ x1360 = 0x13;
-Tbase TZ x1361 = x1360 * x1359;
-Tbase TZ x1362 = x1350 + x1361;
-Tbase TZ x1363 = x1331 + x1362;
-Tbase TZ x1364 = 0x1a;
-Tbase TZ x1365 = x1363 >>> x1364;
-Tbase TZ x1366 = x801 * x448;
-Tbase TZ x1367 = x802 * x445;
-Tbase TZ x1368 = x803 * x442;
-Tbase TZ x1369 = x804 * x439;
-Tbase TZ x1370 = x805 * x436;
-Tbase TZ x1371 = x806 * x433;
-Tbase TZ x1372 = x807 * x430;
-Tbase TZ x1373 = x808 * x427;
-Tbase TZ x1374 = x1372 + x1373;
-Tbase TZ x1375 = x1371 + x1374;
-Tbase TZ x1376 = x1370 + x1375;
-Tbase TZ x1377 = x1369 + x1376;
-Tbase TZ x1378 = x1368 + x1377;
-Tbase TZ x1379 = x1367 + x1378;
-Tbase TZ x1380 = x1366 + x1379;
-Tbase TZ x1381 = x799 * x424;
-Tbase TZ x1382 = x800 * x421;
-Tbase TZ x1383 = x1381 + x1382;
-Tbase TZ x1384 = 0x13;
-Tbase TZ x1385 = x1384 * x1383;
-Tbase TZ x1386 = x1380 + x1385;
-Tbase TZ x1387 = x1365 + x1386;
-Tbase TZ x1388 = 0x19;
-Tbase TZ x1389 = x1387 >>> x1388;
-Tbase TZ x1390 = x800 * x448;
-Tbase TZ x1391 = 0x2;
-Tbase TZ x1392 = x445 * x1391;
-Tbase TZ x1393 = x801 * x1392;
-Tbase TZ x1394 = x802 * x442;
-Tbase TZ x1395 = 0x2;
-Tbase TZ x1396 = x439 * x1395;
-Tbase TZ x1397 = x803 * x1396;
-Tbase TZ x1398 = x804 * x436;
-Tbase TZ x1399 = 0x2;
-Tbase TZ x1400 = x433 * x1399;
-Tbase TZ x1401 = x805 * x1400;
-Tbase TZ x1402 = x806 * x430;
-Tbase TZ x1403 = 0x2;
-Tbase TZ x1404 = x427 * x1403;
-Tbase TZ x1405 = x807 * x1404;
-Tbase TZ x1406 = x808 * x424;
-Tbase TZ x1407 = x1405 + x1406;
-Tbase TZ x1408 = x1402 + x1407;
-Tbase TZ x1409 = x1401 + x1408;
-Tbase TZ x1410 = x1398 + x1409;
-Tbase TZ x1411 = x1397 + x1410;
-Tbase TZ x1412 = x1394 + x1411;
-Tbase TZ x1413 = x1393 + x1412;
-Tbase TZ x1414 = x1390 + x1413;
-Tbase TZ x1415 = 0x2;
-Tbase TZ x1416 = x421 * x1415;
-Tbase TZ x1417 = x799 * x1416;
-Tbase TZ x1418 = 0x13;
-Tbase TZ x1419 = x1418 * x1417;
-Tbase TZ x1420 = x1414 + x1419;
-Tbase TZ x1421 = x1389 + x1420;
-Tbase TZ x1422 = 0x1a;
-Tbase TZ x1423 = x1421 >>> x1422;
-Tbase TZ x1424 = x799 * x448;
-Tbase TZ x1425 = x800 * x445;
-Tbase TZ x1426 = x801 * x442;
-Tbase TZ x1427 = x802 * x439;
-Tbase TZ x1428 = x803 * x436;
-Tbase TZ x1429 = x804 * x433;
-Tbase TZ x1430 = x805 * x430;
-Tbase TZ x1431 = x806 * x427;
-Tbase TZ x1432 = x807 * x424;
-Tbase TZ x1433 = x808 * x421;
-Tbase TZ x1434 = x1432 + x1433;
-Tbase TZ x1435 = x1431 + x1434;
-Tbase TZ x1436 = x1430 + x1435;
-Tbase TZ x1437 = x1429 + x1436;
-Tbase TZ x1438 = x1428 + x1437;
-Tbase TZ x1439 = x1427 + x1438;
-Tbase TZ x1440 = x1426 + x1439;
-Tbase TZ x1441 = x1425 + x1440;
-Tbase TZ x1442 = x1424 + x1441;
-Tbase TZ x1443 = x1423 + x1442;
-Tbase TZ x1444 = 0x19;
-Tbase TZ x1445 = x1443 >>> x1444;
-Tbase TZ x1446 = 0x13;
-Tbase TZ x1447 = x1446 * x1445;
-Tbase TZ x1448 = 0x3ffffff;
-Tbase TZ x1449 = x1189 & x1448;
-Tbase TZ x1450 = x1447 + x1449;
-Tbase TZ x1451 = 0x1a;
-Tbase TZ x1452 = x1450 >>> x1451;
-Tbase TZ x1453 = 0x1ffffff;
-Tbase TZ x1454 = x1213 & x1453;
-Tbase TZ x1455 = x1452 + x1454;
-Tbase TZ x1456 = 0x1ffffff;
-Tbase TZ x1457 = x1443 & x1456;
-Tbase TZ x1458 = 0x3ffffff;
-Tbase TZ x1459 = x1421 & x1458;
-Tbase TZ x1460 = 0x1ffffff;
-Tbase TZ x1461 = x1387 & x1460;
-Tbase TZ x1462 = 0x3ffffff;
-Tbase TZ x1463 = x1363 & x1462;
-Tbase TZ x1464 = 0x1ffffff;
-Tbase TZ x1465 = x1329 & x1464;
-Tbase TZ x1466 = 0x3ffffff;
-Tbase TZ x1467 = x1305 & x1466;
-Tbase TZ x1468 = 0x1ffffff;
-Tbase TZ x1469 = x1271 & x1468;
-Tbase TZ x1470 = 0x19;
-Tbase TZ x1471 = x1455 >>> x1470;
-Tbase TZ x1472 = 0x3ffffff;
-Tbase TZ x1473 = x1247 & x1472;
-Tbase TZ x1474 = x1471 + x1473;
-Tbase TZ x1475 = 0x1ffffff;
-Tbase TZ x1476 = x1455 & x1475;
-Tbase TZ x1477 = 0x3ffffff;
-Tbase TZ x1478 = x1450 & x1477;
-Tbase TZ x1479 = x1137 + x1457;
-Tbase TZ x1480 = x1139 + x1459;
-Tbase TZ x1481 = x1141 + x1461;
-Tbase TZ x1482 = x1143 + x1463;
-Tbase TZ x1483 = x1145 + x1465;
-Tbase TZ x1484 = x1147 + x1467;
-Tbase TZ x1485 = x1149 + x1469;
-Tbase TZ x1486 = x1154 + x1474;
-Tbase TZ x1487 = x1156 + x1476;
-Tbase TZ x1488 = x1158 + x1478;
-Tbase TZ x1489 = x1137 + x1457;
-Tbase TZ x1490 = x1139 + x1459;
-Tbase TZ x1491 = x1141 + x1461;
-Tbase TZ x1492 = x1143 + x1463;
-Tbase TZ x1493 = x1145 + x1465;
-Tbase TZ x1494 = x1147 + x1467;
-Tbase TZ x1495 = x1149 + x1469;
-Tbase TZ x1496 = x1154 + x1474;
-Tbase TZ x1497 = x1156 + x1476;
-Tbase TZ x1498 = x1158 + x1478;
-Tbase TZ x1499 = x1488 * x1498;
-Tbase TZ x1500 = 0x2;
-Tbase TZ x1501 = x1497 * x1500;
-Tbase TZ x1502 = x1479 * x1501;
-Tbase TZ x1503 = x1480 * x1496;
-Tbase TZ x1504 = 0x2;
-Tbase TZ x1505 = x1495 * x1504;
-Tbase TZ x1506 = x1481 * x1505;
-Tbase TZ x1507 = x1482 * x1494;
-Tbase TZ x1508 = 0x2;
-Tbase TZ x1509 = x1493 * x1508;
-Tbase TZ x1510 = x1483 * x1509;
-Tbase TZ x1511 = x1484 * x1492;
-Tbase TZ x1512 = 0x2;
-Tbase TZ x1513 = x1491 * x1512;
-Tbase TZ x1514 = x1485 * x1513;
-Tbase TZ x1515 = x1486 * x1490;
-Tbase TZ x1516 = 0x2;
-Tbase TZ x1517 = x1489 * x1516;
-Tbase TZ x1518 = x1487 * x1517;
-Tbase TZ x1519 = x1515 + x1518;
-Tbase TZ x1520 = x1514 + x1519;
-Tbase TZ x1521 = x1511 + x1520;
-Tbase TZ x1522 = x1510 + x1521;
-Tbase TZ x1523 = x1507 + x1522;
-Tbase TZ x1524 = x1506 + x1523;
-Tbase TZ x1525 = x1503 + x1524;
-Tbase TZ x1526 = x1502 + x1525;
-Tbase TZ x1527 = 0x13;
-Tbase TZ x1528 = x1527 * x1526;
-Tbase TZ x1529 = x1499 + x1528;
-Tbase TZ x1530 = 0x1a;
-Tbase TZ x1531 = x1529 >>> x1530;
-Tbase TZ x1532 = x1487 * x1498;
-Tbase TZ x1533 = x1488 * x1497;
-Tbase TZ x1534 = x1532 + x1533;
-Tbase TZ x1535 = x1479 * x1496;
-Tbase TZ x1536 = x1480 * x1495;
-Tbase TZ x1537 = x1481 * x1494;
-Tbase TZ x1538 = x1482 * x1493;
-Tbase TZ x1539 = x1483 * x1492;
-Tbase TZ x1540 = x1484 * x1491;
-Tbase TZ x1541 = x1485 * x1490;
-Tbase TZ x1542 = x1486 * x1489;
-Tbase TZ x1543 = x1541 + x1542;
-Tbase TZ x1544 = x1540 + x1543;
-Tbase TZ x1545 = x1539 + x1544;
-Tbase TZ x1546 = x1538 + x1545;
-Tbase TZ x1547 = x1537 + x1546;
-Tbase TZ x1548 = x1536 + x1547;
-Tbase TZ x1549 = x1535 + x1548;
-Tbase TZ x1550 = 0x13;
-Tbase TZ x1551 = x1550 * x1549;
-Tbase TZ x1552 = x1534 + x1551;
-Tbase TZ x1553 = x1531 + x1552;
-Tbase TZ x1554 = 0x19;
-Tbase TZ x1555 = x1553 >>> x1554;
-Tbase TZ x1556 = x1486 * x1498;
-Tbase TZ x1557 = 0x2;
-Tbase TZ x1558 = x1497 * x1557;
-Tbase TZ x1559 = x1487 * x1558;
-Tbase TZ x1560 = x1488 * x1496;
-Tbase TZ x1561 = x1559 + x1560;
-Tbase TZ x1562 = x1556 + x1561;
-Tbase TZ x1563 = 0x2;
-Tbase TZ x1564 = x1495 * x1563;
-Tbase TZ x1565 = x1479 * x1564;
-Tbase TZ x1566 = x1480 * x1494;
-Tbase TZ x1567 = 0x2;
-Tbase TZ x1568 = x1493 * x1567;
-Tbase TZ x1569 = x1481 * x1568;
-Tbase TZ x1570 = x1482 * x1492;
-Tbase TZ x1571 = 0x2;
-Tbase TZ x1572 = x1491 * x1571;
-Tbase TZ x1573 = x1483 * x1572;
-Tbase TZ x1574 = x1484 * x1490;
-Tbase TZ x1575 = 0x2;
-Tbase TZ x1576 = x1489 * x1575;
-Tbase TZ x1577 = x1485 * x1576;
-Tbase TZ x1578 = x1574 + x1577;
-Tbase TZ x1579 = x1573 + x1578;
-Tbase TZ x1580 = x1570 + x1579;
-Tbase TZ x1581 = x1569 + x1580;
-Tbase TZ x1582 = x1566 + x1581;
-Tbase TZ x1583 = x1565 + x1582;
-Tbase TZ x1584 = 0x13;
-Tbase TZ x1585 = x1584 * x1583;
-Tbase TZ x1586 = x1562 + x1585;
-Tbase TZ x1587 = x1555 + x1586;
-Tbase TZ x1588 = 0x1a;
-Tbase TZ x1589 = x1587 >>> x1588;
-Tbase TZ x1590 = x1485 * x1498;
-Tbase TZ x1591 = x1486 * x1497;
-Tbase TZ x1592 = x1487 * x1496;
-Tbase TZ x1593 = x1488 * x1495;
-Tbase TZ x1594 = x1592 + x1593;
-Tbase TZ x1595 = x1591 + x1594;
-Tbase TZ x1596 = x1590 + x1595;
-Tbase TZ x1597 = x1479 * x1494;
-Tbase TZ x1598 = x1480 * x1493;
-Tbase TZ x1599 = x1481 * x1492;
-Tbase TZ x1600 = x1482 * x1491;
-Tbase TZ x1601 = x1483 * x1490;
-Tbase TZ x1602 = x1484 * x1489;
-Tbase TZ x1603 = x1601 + x1602;
-Tbase TZ x1604 = x1600 + x1603;
-Tbase TZ x1605 = x1599 + x1604;
-Tbase TZ x1606 = x1598 + x1605;
-Tbase TZ x1607 = x1597 + x1606;
-Tbase TZ x1608 = 0x13;
-Tbase TZ x1609 = x1608 * x1607;
-Tbase TZ x1610 = x1596 + x1609;
-Tbase TZ x1611 = x1589 + x1610;
-Tbase TZ x1612 = 0x19;
-Tbase TZ x1613 = x1611 >>> x1612;
-Tbase TZ x1614 = x1484 * x1498;
-Tbase TZ x1615 = 0x2;
-Tbase TZ x1616 = x1497 * x1615;
-Tbase TZ x1617 = x1485 * x1616;
-Tbase TZ x1618 = x1486 * x1496;
-Tbase TZ x1619 = 0x2;
-Tbase TZ x1620 = x1495 * x1619;
-Tbase TZ x1621 = x1487 * x1620;
-Tbase TZ x1622 = x1488 * x1494;
-Tbase TZ x1623 = x1621 + x1622;
-Tbase TZ x1624 = x1618 + x1623;
-Tbase TZ x1625 = x1617 + x1624;
-Tbase TZ x1626 = x1614 + x1625;
-Tbase TZ x1627 = 0x2;
-Tbase TZ x1628 = x1493 * x1627;
-Tbase TZ x1629 = x1479 * x1628;
-Tbase TZ x1630 = x1480 * x1492;
-Tbase TZ x1631 = 0x2;
-Tbase TZ x1632 = x1491 * x1631;
-Tbase TZ x1633 = x1481 * x1632;
-Tbase TZ x1634 = x1482 * x1490;
-Tbase TZ x1635 = 0x2;
-Tbase TZ x1636 = x1489 * x1635;
-Tbase TZ x1637 = x1483 * x1636;
-Tbase TZ x1638 = x1634 + x1637;
-Tbase TZ x1639 = x1633 + x1638;
-Tbase TZ x1640 = x1630 + x1639;
-Tbase TZ x1641 = x1629 + x1640;
-Tbase TZ x1642 = 0x13;
-Tbase TZ x1643 = x1642 * x1641;
-Tbase TZ x1644 = x1626 + x1643;
-Tbase TZ x1645 = x1613 + x1644;
-Tbase TZ x1646 = 0x1a;
-Tbase TZ x1647 = x1645 >>> x1646;
-Tbase TZ x1648 = x1483 * x1498;
-Tbase TZ x1649 = x1484 * x1497;
-Tbase TZ x1650 = x1485 * x1496;
-Tbase TZ x1651 = x1486 * x1495;
-Tbase TZ x1652 = x1487 * x1494;
-Tbase TZ x1653 = x1488 * x1493;
-Tbase TZ x1654 = x1652 + x1653;
-Tbase TZ x1655 = x1651 + x1654;
-Tbase TZ x1656 = x1650 + x1655;
-Tbase TZ x1657 = x1649 + x1656;
-Tbase TZ x1658 = x1648 + x1657;
-Tbase TZ x1659 = x1479 * x1492;
-Tbase TZ x1660 = x1480 * x1491;
-Tbase TZ x1661 = x1481 * x1490;
-Tbase TZ x1662 = x1482 * x1489;
-Tbase TZ x1663 = x1661 + x1662;
-Tbase TZ x1664 = x1660 + x1663;
-Tbase TZ x1665 = x1659 + x1664;
-Tbase TZ x1666 = 0x13;
-Tbase TZ x1667 = x1666 * x1665;
-Tbase TZ x1668 = x1658 + x1667;
-Tbase TZ x1669 = x1647 + x1668;
-Tbase TZ x1670 = 0x19;
-Tbase TZ x1671 = x1669 >>> x1670;
-Tbase TZ x1672 = x1482 * x1498;
-Tbase TZ x1673 = 0x2;
-Tbase TZ x1674 = x1497 * x1673;
-Tbase TZ x1675 = x1483 * x1674;
-Tbase TZ x1676 = x1484 * x1496;
-Tbase TZ x1677 = 0x2;
-Tbase TZ x1678 = x1495 * x1677;
-Tbase TZ x1679 = x1485 * x1678;
-Tbase TZ x1680 = x1486 * x1494;
-Tbase TZ x1681 = 0x2;
-Tbase TZ x1682 = x1493 * x1681;
-Tbase TZ x1683 = x1487 * x1682;
-Tbase TZ x1684 = x1488 * x1492;
-Tbase TZ x1685 = x1683 + x1684;
-Tbase TZ x1686 = x1680 + x1685;
-Tbase TZ x1687 = x1679 + x1686;
-Tbase TZ x1688 = x1676 + x1687;
-Tbase TZ x1689 = x1675 + x1688;
-Tbase TZ x1690 = x1672 + x1689;
-Tbase TZ x1691 = 0x2;
-Tbase TZ x1692 = x1491 * x1691;
-Tbase TZ x1693 = x1479 * x1692;
-Tbase TZ x1694 = x1480 * x1490;
-Tbase TZ x1695 = 0x2;
-Tbase TZ x1696 = x1489 * x1695;
-Tbase TZ x1697 = x1481 * x1696;
-Tbase TZ x1698 = x1694 + x1697;
-Tbase TZ x1699 = x1693 + x1698;
-Tbase TZ x1700 = 0x13;
-Tbase TZ x1701 = x1700 * x1699;
-Tbase TZ x1702 = x1690 + x1701;
-Tbase TZ x1703 = x1671 + x1702;
-Tbase TZ x1704 = 0x1a;
-Tbase TZ x1705 = x1703 >>> x1704;
-Tbase TZ x1706 = x1481 * x1498;
-Tbase TZ x1707 = x1482 * x1497;
-Tbase TZ x1708 = x1483 * x1496;
-Tbase TZ x1709 = x1484 * x1495;
-Tbase TZ x1710 = x1485 * x1494;
-Tbase TZ x1711 = x1486 * x1493;
-Tbase TZ x1712 = x1487 * x1492;
-Tbase TZ x1713 = x1488 * x1491;
-Tbase TZ x1714 = x1712 + x1713;
-Tbase TZ x1715 = x1711 + x1714;
-Tbase TZ x1716 = x1710 + x1715;
-Tbase TZ x1717 = x1709 + x1716;
-Tbase TZ x1718 = x1708 + x1717;
-Tbase TZ x1719 = x1707 + x1718;
-Tbase TZ x1720 = x1706 + x1719;
-Tbase TZ x1721 = x1479 * x1490;
-Tbase TZ x1722 = x1480 * x1489;
-Tbase TZ x1723 = x1721 + x1722;
-Tbase TZ x1724 = 0x13;
-Tbase TZ x1725 = x1724 * x1723;
-Tbase TZ x1726 = x1720 + x1725;
-Tbase TZ x1727 = x1705 + x1726;
-Tbase TZ x1728 = 0x19;
-Tbase TZ x1729 = x1727 >>> x1728;
-Tbase TZ x1730 = x1480 * x1498;
-Tbase TZ x1731 = 0x2;
-Tbase TZ x1732 = x1497 * x1731;
-Tbase TZ x1733 = x1481 * x1732;
-Tbase TZ x1734 = x1482 * x1496;
-Tbase TZ x1735 = 0x2;
-Tbase TZ x1736 = x1495 * x1735;
-Tbase TZ x1737 = x1483 * x1736;
-Tbase TZ x1738 = x1484 * x1494;
-Tbase TZ x1739 = 0x2;
-Tbase TZ x1740 = x1493 * x1739;
-Tbase TZ x1741 = x1485 * x1740;
-Tbase TZ x1742 = x1486 * x1492;
-Tbase TZ x1743 = 0x2;
-Tbase TZ x1744 = x1491 * x1743;
-Tbase TZ x1745 = x1487 * x1744;
-Tbase TZ x1746 = x1488 * x1490;
-Tbase TZ x1747 = x1745 + x1746;
-Tbase TZ x1748 = x1742 + x1747;
-Tbase TZ x1749 = x1741 + x1748;
-Tbase TZ x1750 = x1738 + x1749;
-Tbase TZ x1751 = x1737 + x1750;
-Tbase TZ x1752 = x1734 + x1751;
-Tbase TZ x1753 = x1733 + x1752;
-Tbase TZ x1754 = x1730 + x1753;
-Tbase TZ x1755 = 0x2;
-Tbase TZ x1756 = x1489 * x1755;
-Tbase TZ x1757 = x1479 * x1756;
-Tbase TZ x1758 = 0x13;
-Tbase TZ x1759 = x1758 * x1757;
-Tbase TZ x1760 = x1754 + x1759;
-Tbase TZ x1761 = x1729 + x1760;
-Tbase TZ x1762 = 0x1a;
-Tbase TZ x1763 = x1761 >>> x1762;
-Tbase TZ x1764 = x1479 * x1498;
-Tbase TZ x1765 = x1480 * x1497;
-Tbase TZ x1766 = x1481 * x1496;
-Tbase TZ x1767 = x1482 * x1495;
-Tbase TZ x1768 = x1483 * x1494;
-Tbase TZ x1769 = x1484 * x1493;
-Tbase TZ x1770 = x1485 * x1492;
-Tbase TZ x1771 = x1486 * x1491;
-Tbase TZ x1772 = x1487 * x1490;
-Tbase TZ x1773 = x1488 * x1489;
-Tbase TZ x1774 = x1772 + x1773;
-Tbase TZ x1775 = x1771 + x1774;
-Tbase TZ x1776 = x1770 + x1775;
-Tbase TZ x1777 = x1769 + x1776;
-Tbase TZ x1778 = x1768 + x1777;
-Tbase TZ x1779 = x1767 + x1778;
-Tbase TZ x1780 = x1766 + x1779;
-Tbase TZ x1781 = x1765 + x1780;
-Tbase TZ x1782 = x1764 + x1781;
-Tbase TZ x1783 = x1763 + x1782;
-Tbase TZ x1784 = 0x19;
-Tbase TZ x1785 = x1783 >>> x1784;
-Tbase TZ x1786 = 0x13;
-Tbase TZ x1787 = x1786 * x1785;
-Tbase TZ x1788 = 0x3ffffff;
-Tbase TZ x1789 = x1529 & x1788;
-Tbase TZ x1790 = x1787 + x1789;
-Tbase TZ x1791 = 0x1a;
-Tbase TZ x1792 = x1790 >>> x1791;
-Tbase TZ x1793 = 0x1ffffff;
-Tbase TZ x1794 = x1553 & x1793;
-Tbase TZ x1795 = x1792 + x1794;
-Tbase TZ x1796 = 0x1ffffff;
-Tbase TZ x1797 = x1783 & x1796;
-Tbase TZ x1798 = 0x3ffffff;
-Tbase TZ x1799 = x1761 & x1798;
-Tbase TZ x1800 = 0x1ffffff;
-Tbase TZ x1801 = x1727 & x1800;
-Tbase TZ x1802 = 0x3ffffff;
-Tbase TZ x1803 = x1703 & x1802;
-Tbase TZ x1804 = 0x1ffffff;
-Tbase TZ x1805 = x1669 & x1804;
-Tbase TZ x1806 = 0x3ffffff;
-Tbase TZ x1807 = x1645 & x1806;
-Tbase TZ x1808 = 0x1ffffff;
-Tbase TZ x1809 = x1611 & x1808;
-Tbase TZ x1810 = 0x19;
-Tbase TZ x1811 = x1795 >>> x1810;
-Tbase TZ x1812 = 0x3ffffff;
-Tbase TZ x1813 = x1587 & x1812;
-Tbase TZ x1814 = x1811 + x1813;
-Tbase TZ x1815 = 0x1ffffff;
-Tbase TZ x1816 = x1795 & x1815;
-Tbase TZ x1817 = 0x3ffffff;
-Tbase TZ x1818 = x1790 & x1817;
-Tbase TZ x1819 = 0x3fffffe;
-Tbase TZ x1820 = x1819 + x1137;
-Tbase TZ x1821 = x1820 - x1457;
-Tbase TZ x1822 = 0x7fffffe;
-Tbase TZ x1823 = x1822 + x1139;
-Tbase TZ x1824 = x1823 - x1459;
-Tbase TZ x1825 = 0x3fffffe;
-Tbase TZ x1826 = x1825 + x1141;
-Tbase TZ x1827 = x1826 - x1461;
-Tbase TZ x1828 = 0x7fffffe;
-Tbase TZ x1829 = x1828 + x1143;
-Tbase TZ x1830 = x1829 - x1463;
-Tbase TZ x1831 = 0x3fffffe;
-Tbase TZ x1832 = x1831 + x1145;
-Tbase TZ x1833 = x1832 - x1465;
-Tbase TZ x1834 = 0x7fffffe;
-Tbase TZ x1835 = x1834 + x1147;
-Tbase TZ x1836 = x1835 - x1467;
-Tbase TZ x1837 = 0x3fffffe;
-Tbase TZ x1838 = x1837 + x1149;
-Tbase TZ x1839 = x1838 - x1469;
-Tbase TZ x1840 = 0x7fffffe;
-Tbase TZ x1841 = x1840 + x1154;
-Tbase TZ x1842 = x1841 - x1474;
-Tbase TZ x1843 = 0x3fffffe;
-Tbase TZ x1844 = x1843 + x1156;
-Tbase TZ x1845 = x1844 - x1476;
-Tbase TZ x1846 = 0x7ffffda;
-Tbase TZ x1847 = x1846 + x1158;
-Tbase TZ x1848 = x1847 - x1478;
-Tbase TZ x1849 = 0x3fffffe;
-Tbase TZ x1850 = x1849 + x1137;
-Tbase TZ x1851 = x1850 - x1457;
-Tbase TZ x1852 = 0x7fffffe;
-Tbase TZ x1853 = x1852 + x1139;
-Tbase TZ x1854 = x1853 - x1459;
-Tbase TZ x1855 = 0x3fffffe;
-Tbase TZ x1856 = x1855 + x1141;
-Tbase TZ x1857 = x1856 - x1461;
-Tbase TZ x1858 = 0x7fffffe;
-Tbase TZ x1859 = x1858 + x1143;
-Tbase TZ x1860 = x1859 - x1463;
-Tbase TZ x1861 = 0x3fffffe;
-Tbase TZ x1862 = x1861 + x1145;
-Tbase TZ x1863 = x1862 - x1465;
-Tbase TZ x1864 = 0x7fffffe;
-Tbase TZ x1865 = x1864 + x1147;
-Tbase TZ x1866 = x1865 - x1467;
-Tbase TZ x1867 = 0x3fffffe;
-Tbase TZ x1868 = x1867 + x1149;
-Tbase TZ x1869 = x1868 - x1469;
-Tbase TZ x1870 = 0x7fffffe;
-Tbase TZ x1871 = x1870 + x1154;
-Tbase TZ x1872 = x1871 - x1474;
-Tbase TZ x1873 = 0x3fffffe;
-Tbase TZ x1874 = x1873 + x1156;
-Tbase TZ x1875 = x1874 - x1476;
-Tbase TZ x1876 = 0x7ffffda;
-Tbase TZ x1877 = x1876 + x1158;
-Tbase TZ x1878 = x1877 - x1478;
-Tbase TZ x1879 = x1848 * x1878;
-Tbase TZ x1880 = 0x2;
-Tbase TZ x1881 = x1875 * x1880;
-Tbase TZ x1882 = x1821 * x1881;
-Tbase TZ x1883 = x1824 * x1872;
-Tbase TZ x1884 = 0x2;
-Tbase TZ x1885 = x1869 * x1884;
-Tbase TZ x1886 = x1827 * x1885;
-Tbase TZ x1887 = x1830 * x1866;
-Tbase TZ x1888 = 0x2;
-Tbase TZ x1889 = x1863 * x1888;
-Tbase TZ x1890 = x1833 * x1889;
-Tbase TZ x1891 = x1836 * x1860;
-Tbase TZ x1892 = 0x2;
-Tbase TZ x1893 = x1857 * x1892;
-Tbase TZ x1894 = x1839 * x1893;
-Tbase TZ x1895 = x1842 * x1854;
-Tbase TZ x1896 = 0x2;
-Tbase TZ x1897 = x1851 * x1896;
-Tbase TZ x1898 = x1845 * x1897;
-Tbase TZ x1899 = x1895 + x1898;
-Tbase TZ x1900 = x1894 + x1899;
-Tbase TZ x1901 = x1891 + x1900;
-Tbase TZ x1902 = x1890 + x1901;
-Tbase TZ x1903 = x1887 + x1902;
-Tbase TZ x1904 = x1886 + x1903;
-Tbase TZ x1905 = x1883 + x1904;
-Tbase TZ x1906 = x1882 + x1905;
-Tbase TZ x1907 = 0x13;
-Tbase TZ x1908 = x1907 * x1906;
-Tbase TZ x1909 = x1879 + x1908;
-Tbase TZ x1910 = 0x1a;
-Tbase TZ x1911 = x1909 >>> x1910;
-Tbase TZ x1912 = x1845 * x1878;
-Tbase TZ x1913 = x1848 * x1875;
-Tbase TZ x1914 = x1912 + x1913;
-Tbase TZ x1915 = x1821 * x1872;
-Tbase TZ x1916 = x1824 * x1869;
-Tbase TZ x1917 = x1827 * x1866;
-Tbase TZ x1918 = x1830 * x1863;
-Tbase TZ x1919 = x1833 * x1860;
-Tbase TZ x1920 = x1836 * x1857;
-Tbase TZ x1921 = x1839 * x1854;
-Tbase TZ x1922 = x1842 * x1851;
-Tbase TZ x1923 = x1921 + x1922;
-Tbase TZ x1924 = x1920 + x1923;
-Tbase TZ x1925 = x1919 + x1924;
-Tbase TZ x1926 = x1918 + x1925;
-Tbase TZ x1927 = x1917 + x1926;
-Tbase TZ x1928 = x1916 + x1927;
-Tbase TZ x1929 = x1915 + x1928;
-Tbase TZ x1930 = 0x13;
-Tbase TZ x1931 = x1930 * x1929;
-Tbase TZ x1932 = x1914 + x1931;
-Tbase TZ x1933 = x1911 + x1932;
-Tbase TZ x1934 = 0x19;
-Tbase TZ x1935 = x1933 >>> x1934;
-Tbase TZ x1936 = x1842 * x1878;
-Tbase TZ x1937 = 0x2;
-Tbase TZ x1938 = x1875 * x1937;
-Tbase TZ x1939 = x1845 * x1938;
-Tbase TZ x1940 = x1848 * x1872;
-Tbase TZ x1941 = x1939 + x1940;
-Tbase TZ x1942 = x1936 + x1941;
-Tbase TZ x1943 = 0x2;
-Tbase TZ x1944 = x1869 * x1943;
-Tbase TZ x1945 = x1821 * x1944;
-Tbase TZ x1946 = x1824 * x1866;
-Tbase TZ x1947 = 0x2;
-Tbase TZ x1948 = x1863 * x1947;
-Tbase TZ x1949 = x1827 * x1948;
-Tbase TZ x1950 = x1830 * x1860;
-Tbase TZ x1951 = 0x2;
-Tbase TZ x1952 = x1857 * x1951;
-Tbase TZ x1953 = x1833 * x1952;
-Tbase TZ x1954 = x1836 * x1854;
-Tbase TZ x1955 = 0x2;
-Tbase TZ x1956 = x1851 * x1955;
-Tbase TZ x1957 = x1839 * x1956;
-Tbase TZ x1958 = x1954 + x1957;
-Tbase TZ x1959 = x1953 + x1958;
-Tbase TZ x1960 = x1950 + x1959;
-Tbase TZ x1961 = x1949 + x1960;
-Tbase TZ x1962 = x1946 + x1961;
-Tbase TZ x1963 = x1945 + x1962;
-Tbase TZ x1964 = 0x13;
-Tbase TZ x1965 = x1964 * x1963;
-Tbase TZ x1966 = x1942 + x1965;
-Tbase TZ x1967 = x1935 + x1966;
-Tbase TZ x1968 = 0x1a;
-Tbase TZ x1969 = x1967 >>> x1968;
-Tbase TZ x1970 = x1839 * x1878;
-Tbase TZ x1971 = x1842 * x1875;
-Tbase TZ x1972 = x1845 * x1872;
-Tbase TZ x1973 = x1848 * x1869;
-Tbase TZ x1974 = x1972 + x1973;
-Tbase TZ x1975 = x1971 + x1974;
-Tbase TZ x1976 = x1970 + x1975;
-Tbase TZ x1977 = x1821 * x1866;
-Tbase TZ x1978 = x1824 * x1863;
-Tbase TZ x1979 = x1827 * x1860;
-Tbase TZ x1980 = x1830 * x1857;
-Tbase TZ x1981 = x1833 * x1854;
-Tbase TZ x1982 = x1836 * x1851;
-Tbase TZ x1983 = x1981 + x1982;
-Tbase TZ x1984 = x1980 + x1983;
-Tbase TZ x1985 = x1979 + x1984;
-Tbase TZ x1986 = x1978 + x1985;
-Tbase TZ x1987 = x1977 + x1986;
-Tbase TZ x1988 = 0x13;
-Tbase TZ x1989 = x1988 * x1987;
-Tbase TZ x1990 = x1976 + x1989;
-Tbase TZ x1991 = x1969 + x1990;
-Tbase TZ x1992 = 0x19;
-Tbase TZ x1993 = x1991 >>> x1992;
-Tbase TZ x1994 = x1836 * x1878;
-Tbase TZ x1995 = 0x2;
-Tbase TZ x1996 = x1875 * x1995;
-Tbase TZ x1997 = x1839 * x1996;
-Tbase TZ x1998 = x1842 * x1872;
-Tbase TZ x1999 = 0x2;
-Tbase TZ x2000 = x1869 * x1999;
-Tbase TZ x2001 = x1845 * x2000;
-Tbase TZ x2002 = x1848 * x1866;
-Tbase TZ x2003 = x2001 + x2002;
-Tbase TZ x2004 = x1998 + x2003;
-Tbase TZ x2005 = x1997 + x2004;
-Tbase TZ x2006 = x1994 + x2005;
-Tbase TZ x2007 = 0x2;
-Tbase TZ x2008 = x1863 * x2007;
-Tbase TZ x2009 = x1821 * x2008;
-Tbase TZ x2010 = x1824 * x1860;
-Tbase TZ x2011 = 0x2;
-Tbase TZ x2012 = x1857 * x2011;
-Tbase TZ x2013 = x1827 * x2012;
-Tbase TZ x2014 = x1830 * x1854;
-Tbase TZ x2015 = 0x2;
-Tbase TZ x2016 = x1851 * x2015;
-Tbase TZ x2017 = x1833 * x2016;
-Tbase TZ x2018 = x2014 + x2017;
-Tbase TZ x2019 = x2013 + x2018;
-Tbase TZ x2020 = x2010 + x2019;
-Tbase TZ x2021 = x2009 + x2020;
-Tbase TZ x2022 = 0x13;
-Tbase TZ x2023 = x2022 * x2021;
-Tbase TZ x2024 = x2006 + x2023;
-Tbase TZ x2025 = x1993 + x2024;
-Tbase TZ x2026 = 0x1a;
-Tbase TZ x2027 = x2025 >>> x2026;
-Tbase TZ x2028 = x1833 * x1878;
-Tbase TZ x2029 = x1836 * x1875;
-Tbase TZ x2030 = x1839 * x1872;
-Tbase TZ x2031 = x1842 * x1869;
-Tbase TZ x2032 = x1845 * x1866;
-Tbase TZ x2033 = x1848 * x1863;
-Tbase TZ x2034 = x2032 + x2033;
-Tbase TZ x2035 = x2031 + x2034;
-Tbase TZ x2036 = x2030 + x2035;
-Tbase TZ x2037 = x2029 + x2036;
-Tbase TZ x2038 = x2028 + x2037;
-Tbase TZ x2039 = x1821 * x1860;
-Tbase TZ x2040 = x1824 * x1857;
-Tbase TZ x2041 = x1827 * x1854;
-Tbase TZ x2042 = x1830 * x1851;
-Tbase TZ x2043 = x2041 + x2042;
-Tbase TZ x2044 = x2040 + x2043;
-Tbase TZ x2045 = x2039 + x2044;
-Tbase TZ x2046 = 0x13;
-Tbase TZ x2047 = x2046 * x2045;
-Tbase TZ x2048 = x2038 + x2047;
-Tbase TZ x2049 = x2027 + x2048;
-Tbase TZ x2050 = 0x19;
-Tbase TZ x2051 = x2049 >>> x2050;
-Tbase TZ x2052 = x1830 * x1878;
-Tbase TZ x2053 = 0x2;
-Tbase TZ x2054 = x1875 * x2053;
-Tbase TZ x2055 = x1833 * x2054;
-Tbase TZ x2056 = x1836 * x1872;
-Tbase TZ x2057 = 0x2;
-Tbase TZ x2058 = x1869 * x2057;
-Tbase TZ x2059 = x1839 * x2058;
-Tbase TZ x2060 = x1842 * x1866;
-Tbase TZ x2061 = 0x2;
-Tbase TZ x2062 = x1863 * x2061;
-Tbase TZ x2063 = x1845 * x2062;
-Tbase TZ x2064 = x1848 * x1860;
-Tbase TZ x2065 = x2063 + x2064;
-Tbase TZ x2066 = x2060 + x2065;
-Tbase TZ x2067 = x2059 + x2066;
-Tbase TZ x2068 = x2056 + x2067;
-Tbase TZ x2069 = x2055 + x2068;
-Tbase TZ x2070 = x2052 + x2069;
-Tbase TZ x2071 = 0x2;
-Tbase TZ x2072 = x1857 * x2071;
-Tbase TZ x2073 = x1821 * x2072;
-Tbase TZ x2074 = x1824 * x1854;
-Tbase TZ x2075 = 0x2;
-Tbase TZ x2076 = x1851 * x2075;
-Tbase TZ x2077 = x1827 * x2076;
-Tbase TZ x2078 = x2074 + x2077;
-Tbase TZ x2079 = x2073 + x2078;
-Tbase TZ x2080 = 0x13;
-Tbase TZ x2081 = x2080 * x2079;
-Tbase TZ x2082 = x2070 + x2081;
-Tbase TZ x2083 = x2051 + x2082;
-Tbase TZ x2084 = 0x1a;
-Tbase TZ x2085 = x2083 >>> x2084;
-Tbase TZ x2086 = x1827 * x1878;
-Tbase TZ x2087 = x1830 * x1875;
-Tbase TZ x2088 = x1833 * x1872;
-Tbase TZ x2089 = x1836 * x1869;
-Tbase TZ x2090 = x1839 * x1866;
-Tbase TZ x2091 = x1842 * x1863;
-Tbase TZ x2092 = x1845 * x1860;
-Tbase TZ x2093 = x1848 * x1857;
-Tbase TZ x2094 = x2092 + x2093;
-Tbase TZ x2095 = x2091 + x2094;
-Tbase TZ x2096 = x2090 + x2095;
-Tbase TZ x2097 = x2089 + x2096;
-Tbase TZ x2098 = x2088 + x2097;
-Tbase TZ x2099 = x2087 + x2098;
-Tbase TZ x2100 = x2086 + x2099;
-Tbase TZ x2101 = x1821 * x1854;
-Tbase TZ x2102 = x1824 * x1851;
-Tbase TZ x2103 = x2101 + x2102;
-Tbase TZ x2104 = 0x13;
-Tbase TZ x2105 = x2104 * x2103;
-Tbase TZ x2106 = x2100 + x2105;
-Tbase TZ x2107 = x2085 + x2106;
-Tbase TZ x2108 = 0x19;
-Tbase TZ x2109 = x2107 >>> x2108;
-Tbase TZ x2110 = x1824 * x1878;
-Tbase TZ x2111 = 0x2;
-Tbase TZ x2112 = x1875 * x2111;
-Tbase TZ x2113 = x1827 * x2112;
-Tbase TZ x2114 = x1830 * x1872;
-Tbase TZ x2115 = 0x2;
-Tbase TZ x2116 = x1869 * x2115;
-Tbase TZ x2117 = x1833 * x2116;
-Tbase TZ x2118 = x1836 * x1866;
-Tbase TZ x2119 = 0x2;
-Tbase TZ x2120 = x1863 * x2119;
-Tbase TZ x2121 = x1839 * x2120;
-Tbase TZ x2122 = x1842 * x1860;
-Tbase TZ x2123 = 0x2;
-Tbase TZ x2124 = x1857 * x2123;
-Tbase TZ x2125 = x1845 * x2124;
-Tbase TZ x2126 = x1848 * x1854;
-Tbase TZ x2127 = x2125 + x2126;
-Tbase TZ x2128 = x2122 + x2127;
-Tbase TZ x2129 = x2121 + x2128;
-Tbase TZ x2130 = x2118 + x2129;
-Tbase TZ x2131 = x2117 + x2130;
-Tbase TZ x2132 = x2114 + x2131;
-Tbase TZ x2133 = x2113 + x2132;
-Tbase TZ x2134 = x2110 + x2133;
-Tbase TZ x2135 = 0x2;
-Tbase TZ x2136 = x1851 * x2135;
-Tbase TZ x2137 = x1821 * x2136;
-Tbase TZ x2138 = 0x13;
-Tbase TZ x2139 = x2138 * x2137;
-Tbase TZ x2140 = x2134 + x2139;
-Tbase TZ x2141 = x2109 + x2140;
-Tbase TZ x2142 = 0x1a;
-Tbase TZ x2143 = x2141 >>> x2142;
-Tbase TZ x2144 = x1821 * x1878;
-Tbase TZ x2145 = x1824 * x1875;
-Tbase TZ x2146 = x1827 * x1872;
-Tbase TZ x2147 = x1830 * x1869;
-Tbase TZ x2148 = x1833 * x1866;
-Tbase TZ x2149 = x1836 * x1863;
-Tbase TZ x2150 = x1839 * x1860;
-Tbase TZ x2151 = x1842 * x1857;
-Tbase TZ x2152 = x1845 * x1854;
-Tbase TZ x2153 = x1848 * x1851;
-Tbase TZ x2154 = x2152 + x2153;
-Tbase TZ x2155 = x2151 + x2154;
-Tbase TZ x2156 = x2150 + x2155;
-Tbase TZ x2157 = x2149 + x2156;
-Tbase TZ x2158 = x2148 + x2157;
-Tbase TZ x2159 = x2147 + x2158;
-Tbase TZ x2160 = x2146 + x2159;
-Tbase TZ x2161 = x2145 + x2160;
-Tbase TZ x2162 = x2144 + x2161;
-Tbase TZ x2163 = x2143 + x2162;
-Tbase TZ x2164 = 0x19;
-Tbase TZ x2165 = x2163 >>> x2164;
-Tbase TZ x2166 = 0x13;
-Tbase TZ x2167 = x2166 * x2165;
-Tbase TZ x2168 = 0x3ffffff;
-Tbase TZ x2169 = x1909 & x2168;
-Tbase TZ x2170 = x2167 + x2169;
-Tbase TZ x2171 = 0x1a;
-Tbase TZ x2172 = x2170 >>> x2171;
-Tbase TZ x2173 = 0x1ffffff;
-Tbase TZ x2174 = x1933 & x2173;
-Tbase TZ x2175 = x2172 + x2174;
-Tbase TZ x2176 = 0x1ffffff;
-Tbase TZ x2177 = x2163 & x2176;
-Tbase TZ x2178 = 0x3ffffff;
-Tbase TZ x2179 = x2141 & x2178;
-Tbase TZ x2180 = 0x1ffffff;
-Tbase TZ x2181 = x2107 & x2180;
-Tbase TZ x2182 = 0x3ffffff;
-Tbase TZ x2183 = x2083 & x2182;
-Tbase TZ x2184 = 0x1ffffff;
-Tbase TZ x2185 = x2049 & x2184;
-Tbase TZ x2186 = 0x3ffffff;
-Tbase TZ x2187 = x2025 & x2186;
-Tbase TZ x2188 = 0x1ffffff;
-Tbase TZ x2189 = x1991 & x2188;
-Tbase TZ x2190 = 0x19;
-Tbase TZ x2191 = x2175 >>> x2190;
-Tbase TZ x2192 = 0x3ffffff;
-Tbase TZ x2193 = x1967 & x2192;
-Tbase TZ x2194 = x2191 + x2193;
-Tbase TZ x2195 = 0x1ffffff;
-Tbase TZ x2196 = x2175 & x2195;
-Tbase TZ x2197 = 0x3ffffff;
-Tbase TZ x2198 = x2170 & x2197;
-Tbase TZ x2199 = x48 * x2198;
-Tbase TZ x2200 = 0x2;
-Tbase TZ x2201 = x2196 * x2200;
-Tbase TZ x2202 = x39 * x2201;
-Tbase TZ x2203 = x40 * x2194;
-Tbase TZ x2204 = 0x2;
-Tbase TZ x2205 = x2189 * x2204;
-Tbase TZ x2206 = x41 * x2205;
-Tbase TZ x2207 = x42 * x2187;
-Tbase TZ x2208 = 0x2;
-Tbase TZ x2209 = x2185 * x2208;
-Tbase TZ x2210 = x43 * x2209;
-Tbase TZ x2211 = x44 * x2183;
-Tbase TZ x2212 = 0x2;
-Tbase TZ x2213 = x2181 * x2212;
-Tbase TZ x2214 = x45 * x2213;
-Tbase TZ x2215 = x46 * x2179;
-Tbase TZ x2216 = 0x2;
-Tbase TZ x2217 = x2177 * x2216;
-Tbase TZ x2218 = x47 * x2217;
-Tbase TZ x2219 = x2215 + x2218;
-Tbase TZ x2220 = x2214 + x2219;
-Tbase TZ x2221 = x2211 + x2220;
-Tbase TZ x2222 = x2210 + x2221;
-Tbase TZ x2223 = x2207 + x2222;
-Tbase TZ x2224 = x2206 + x2223;
-Tbase TZ x2225 = x2203 + x2224;
-Tbase TZ x2226 = x2202 + x2225;
-Tbase TZ x2227 = 0x13;
-Tbase TZ x2228 = x2227 * x2226;
-Tbase TZ x2229 = x2199 + x2228;
-Tbase TZ x2230 = 0x1a;
-Tbase TZ x2231 = x2229 >>> x2230;
-Tbase TZ x2232 = x47 * x2198;
-Tbase TZ x2233 = x48 * x2196;
-Tbase TZ x2234 = x2232 + x2233;
-Tbase TZ x2235 = x39 * x2194;
-Tbase TZ x2236 = x40 * x2189;
-Tbase TZ x2237 = x41 * x2187;
-Tbase TZ x2238 = x42 * x2185;
-Tbase TZ x2239 = x43 * x2183;
-Tbase TZ x2240 = x44 * x2181;
-Tbase TZ x2241 = x45 * x2179;
-Tbase TZ x2242 = x46 * x2177;
-Tbase TZ x2243 = x2241 + x2242;
-Tbase TZ x2244 = x2240 + x2243;
-Tbase TZ x2245 = x2239 + x2244;
-Tbase TZ x2246 = x2238 + x2245;
-Tbase TZ x2247 = x2237 + x2246;
-Tbase TZ x2248 = x2236 + x2247;
-Tbase TZ x2249 = x2235 + x2248;
-Tbase TZ x2250 = 0x13;
-Tbase TZ x2251 = x2250 * x2249;
-Tbase TZ x2252 = x2234 + x2251;
-Tbase TZ x2253 = x2231 + x2252;
-Tbase TZ x2254 = 0x19;
-Tbase TZ x2255 = x2253 >>> x2254;
-Tbase TZ x2256 = x46 * x2198;
-Tbase TZ x2257 = 0x2;
-Tbase TZ x2258 = x2196 * x2257;
-Tbase TZ x2259 = x47 * x2258;
-Tbase TZ x2260 = x48 * x2194;
-Tbase TZ x2261 = x2259 + x2260;
-Tbase TZ x2262 = x2256 + x2261;
-Tbase TZ x2263 = 0x2;
-Tbase TZ x2264 = x2189 * x2263;
-Tbase TZ x2265 = x39 * x2264;
-Tbase TZ x2266 = x40 * x2187;
-Tbase TZ x2267 = 0x2;
-Tbase TZ x2268 = x2185 * x2267;
-Tbase TZ x2269 = x41 * x2268;
-Tbase TZ x2270 = x42 * x2183;
-Tbase TZ x2271 = 0x2;
-Tbase TZ x2272 = x2181 * x2271;
-Tbase TZ x2273 = x43 * x2272;
-Tbase TZ x2274 = x44 * x2179;
-Tbase TZ x2275 = 0x2;
-Tbase TZ x2276 = x2177 * x2275;
-Tbase TZ x2277 = x45 * x2276;
-Tbase TZ x2278 = x2274 + x2277;
-Tbase TZ x2279 = x2273 + x2278;
-Tbase TZ x2280 = x2270 + x2279;
-Tbase TZ x2281 = x2269 + x2280;
-Tbase TZ x2282 = x2266 + x2281;
-Tbase TZ x2283 = x2265 + x2282;
-Tbase TZ x2284 = 0x13;
-Tbase TZ x2285 = x2284 * x2283;
-Tbase TZ x2286 = x2262 + x2285;
-Tbase TZ x2287 = x2255 + x2286;
-Tbase TZ x2288 = 0x1a;
-Tbase TZ x2289 = x2287 >>> x2288;
-Tbase TZ x2290 = x45 * x2198;
-Tbase TZ x2291 = x46 * x2196;
-Tbase TZ x2292 = x47 * x2194;
-Tbase TZ x2293 = x48 * x2189;
-Tbase TZ x2294 = x2292 + x2293;
-Tbase TZ x2295 = x2291 + x2294;
-Tbase TZ x2296 = x2290 + x2295;
-Tbase TZ x2297 = x39 * x2187;
-Tbase TZ x2298 = x40 * x2185;
-Tbase TZ x2299 = x41 * x2183;
-Tbase TZ x2300 = x42 * x2181;
-Tbase TZ x2301 = x43 * x2179;
-Tbase TZ x2302 = x44 * x2177;
-Tbase TZ x2303 = x2301 + x2302;
-Tbase TZ x2304 = x2300 + x2303;
-Tbase TZ x2305 = x2299 + x2304;
-Tbase TZ x2306 = x2298 + x2305;
-Tbase TZ x2307 = x2297 + x2306;
-Tbase TZ x2308 = 0x13;
-Tbase TZ x2309 = x2308 * x2307;
-Tbase TZ x2310 = x2296 + x2309;
-Tbase TZ x2311 = x2289 + x2310;
-Tbase TZ x2312 = 0x19;
-Tbase TZ x2313 = x2311 >>> x2312;
-Tbase TZ x2314 = x44 * x2198;
-Tbase TZ x2315 = 0x2;
-Tbase TZ x2316 = x2196 * x2315;
-Tbase TZ x2317 = x45 * x2316;
-Tbase TZ x2318 = x46 * x2194;
-Tbase TZ x2319 = 0x2;
-Tbase TZ x2320 = x2189 * x2319;
-Tbase TZ x2321 = x47 * x2320;
-Tbase TZ x2322 = x48 * x2187;
-Tbase TZ x2323 = x2321 + x2322;
-Tbase TZ x2324 = x2318 + x2323;
-Tbase TZ x2325 = x2317 + x2324;
-Tbase TZ x2326 = x2314 + x2325;
-Tbase TZ x2327 = 0x2;
-Tbase TZ x2328 = x2185 * x2327;
-Tbase TZ x2329 = x39 * x2328;
-Tbase TZ x2330 = x40 * x2183;
-Tbase TZ x2331 = 0x2;
-Tbase TZ x2332 = x2181 * x2331;
-Tbase TZ x2333 = x41 * x2332;
-Tbase TZ x2334 = x42 * x2179;
-Tbase TZ x2335 = 0x2;
-Tbase TZ x2336 = x2177 * x2335;
-Tbase TZ x2337 = x43 * x2336;
-Tbase TZ x2338 = x2334 + x2337;
-Tbase TZ x2339 = x2333 + x2338;
-Tbase TZ x2340 = x2330 + x2339;
-Tbase TZ x2341 = x2329 + x2340;
-Tbase TZ x2342 = 0x13;
-Tbase TZ x2343 = x2342 * x2341;
-Tbase TZ x2344 = x2326 + x2343;
-Tbase TZ x2345 = x2313 + x2344;
-Tbase TZ x2346 = 0x1a;
-Tbase TZ x2347 = x2345 >>> x2346;
-Tbase TZ x2348 = x43 * x2198;
-Tbase TZ x2349 = x44 * x2196;
-Tbase TZ x2350 = x45 * x2194;
-Tbase TZ x2351 = x46 * x2189;
-Tbase TZ x2352 = x47 * x2187;
-Tbase TZ x2353 = x48 * x2185;
-Tbase TZ x2354 = x2352 + x2353;
-Tbase TZ x2355 = x2351 + x2354;
-Tbase TZ x2356 = x2350 + x2355;
-Tbase TZ x2357 = x2349 + x2356;
-Tbase TZ x2358 = x2348 + x2357;
-Tbase TZ x2359 = x39 * x2183;
-Tbase TZ x2360 = x40 * x2181;
-Tbase TZ x2361 = x41 * x2179;
-Tbase TZ x2362 = x42 * x2177;
-Tbase TZ x2363 = x2361 + x2362;
-Tbase TZ x2364 = x2360 + x2363;
-Tbase TZ x2365 = x2359 + x2364;
-Tbase TZ x2366 = 0x13;
-Tbase TZ x2367 = x2366 * x2365;
-Tbase TZ x2368 = x2358 + x2367;
-Tbase TZ x2369 = x2347 + x2368;
-Tbase TZ x2370 = 0x19;
-Tbase TZ x2371 = x2369 >>> x2370;
-Tbase TZ x2372 = x42 * x2198;
-Tbase TZ x2373 = 0x2;
-Tbase TZ x2374 = x2196 * x2373;
-Tbase TZ x2375 = x43 * x2374;
-Tbase TZ x2376 = x44 * x2194;
-Tbase TZ x2377 = 0x2;
-Tbase TZ x2378 = x2189 * x2377;
-Tbase TZ x2379 = x45 * x2378;
-Tbase TZ x2380 = x46 * x2187;
-Tbase TZ x2381 = 0x2;
-Tbase TZ x2382 = x2185 * x2381;
-Tbase TZ x2383 = x47 * x2382;
-Tbase TZ x2384 = x48 * x2183;
-Tbase TZ x2385 = x2383 + x2384;
-Tbase TZ x2386 = x2380 + x2385;
-Tbase TZ x2387 = x2379 + x2386;
-Tbase TZ x2388 = x2376 + x2387;
-Tbase TZ x2389 = x2375 + x2388;
-Tbase TZ x2390 = x2372 + x2389;
-Tbase TZ x2391 = 0x2;
-Tbase TZ x2392 = x2181 * x2391;
-Tbase TZ x2393 = x39 * x2392;
-Tbase TZ x2394 = x40 * x2179;
-Tbase TZ x2395 = 0x2;
-Tbase TZ x2396 = x2177 * x2395;
-Tbase TZ x2397 = x41 * x2396;
-Tbase TZ x2398 = x2394 + x2397;
-Tbase TZ x2399 = x2393 + x2398;
-Tbase TZ x2400 = 0x13;
-Tbase TZ x2401 = x2400 * x2399;
-Tbase TZ x2402 = x2390 + x2401;
-Tbase TZ x2403 = x2371 + x2402;
-Tbase TZ x2404 = 0x1a;
-Tbase TZ x2405 = x2403 >>> x2404;
-Tbase TZ x2406 = x41 * x2198;
-Tbase TZ x2407 = x42 * x2196;
-Tbase TZ x2408 = x43 * x2194;
-Tbase TZ x2409 = x44 * x2189;
-Tbase TZ x2410 = x45 * x2187;
-Tbase TZ x2411 = x46 * x2185;
-Tbase TZ x2412 = x47 * x2183;
-Tbase TZ x2413 = x48 * x2181;
-Tbase TZ x2414 = x2412 + x2413;
-Tbase TZ x2415 = x2411 + x2414;
-Tbase TZ x2416 = x2410 + x2415;
-Tbase TZ x2417 = x2409 + x2416;
-Tbase TZ x2418 = x2408 + x2417;
-Tbase TZ x2419 = x2407 + x2418;
-Tbase TZ x2420 = x2406 + x2419;
-Tbase TZ x2421 = x39 * x2179;
-Tbase TZ x2422 = x40 * x2177;
-Tbase TZ x2423 = x2421 + x2422;
-Tbase TZ x2424 = 0x13;
-Tbase TZ x2425 = x2424 * x2423;
-Tbase TZ x2426 = x2420 + x2425;
-Tbase TZ x2427 = x2405 + x2426;
-Tbase TZ x2428 = 0x19;
-Tbase TZ x2429 = x2427 >>> x2428;
-Tbase TZ x2430 = x40 * x2198;
-Tbase TZ x2431 = 0x2;
-Tbase TZ x2432 = x2196 * x2431;
-Tbase TZ x2433 = x41 * x2432;
-Tbase TZ x2434 = x42 * x2194;
-Tbase TZ x2435 = 0x2;
-Tbase TZ x2436 = x2189 * x2435;
-Tbase TZ x2437 = x43 * x2436;
-Tbase TZ x2438 = x44 * x2187;
-Tbase TZ x2439 = 0x2;
-Tbase TZ x2440 = x2185 * x2439;
-Tbase TZ x2441 = x45 * x2440;
-Tbase TZ x2442 = x46 * x2183;
-Tbase TZ x2443 = 0x2;
-Tbase TZ x2444 = x2181 * x2443;
-Tbase TZ x2445 = x47 * x2444;
-Tbase TZ x2446 = x48 * x2179;
-Tbase TZ x2447 = x2445 + x2446;
-Tbase TZ x2448 = x2442 + x2447;
-Tbase TZ x2449 = x2441 + x2448;
-Tbase TZ x2450 = x2438 + x2449;
-Tbase TZ x2451 = x2437 + x2450;
-Tbase TZ x2452 = x2434 + x2451;
-Tbase TZ x2453 = x2433 + x2452;
-Tbase TZ x2454 = x2430 + x2453;
-Tbase TZ x2455 = 0x2;
-Tbase TZ x2456 = x2177 * x2455;
-Tbase TZ x2457 = x39 * x2456;
-Tbase TZ x2458 = 0x13;
-Tbase TZ x2459 = x2458 * x2457;
-Tbase TZ x2460 = x2454 + x2459;
-Tbase TZ x2461 = x2429 + x2460;
-Tbase TZ x2462 = 0x1a;
-Tbase TZ x2463 = x2461 >>> x2462;
-Tbase TZ x2464 = x39 * x2198;
-Tbase TZ x2465 = x40 * x2196;
-Tbase TZ x2466 = x41 * x2194;
-Tbase TZ x2467 = x42 * x2189;
-Tbase TZ x2468 = x43 * x2187;
-Tbase TZ x2469 = x44 * x2185;
-Tbase TZ x2470 = x45 * x2183;
-Tbase TZ x2471 = x46 * x2181;
-Tbase TZ x2472 = x47 * x2179;
-Tbase TZ x2473 = x48 * x2177;
-Tbase TZ x2474 = x2472 + x2473;
-Tbase TZ x2475 = x2471 + x2474;
-Tbase TZ x2476 = x2470 + x2475;
-Tbase TZ x2477 = x2469 + x2476;
-Tbase TZ x2478 = x2468 + x2477;
-Tbase TZ x2479 = x2467 + x2478;
-Tbase TZ x2480 = x2466 + x2479;
-Tbase TZ x2481 = x2465 + x2480;
-Tbase TZ x2482 = x2464 + x2481;
-Tbase TZ x2483 = x2463 + x2482;
-Tbase TZ x2484 = 0x19;
-Tbase TZ x2485 = x2483 >>> x2484;
-Tbase TZ x2486 = 0x13;
-Tbase TZ x2487 = x2486 * x2485;
-Tbase TZ x2488 = 0x3ffffff;
-Tbase TZ x2489 = x2229 & x2488;
-Tbase TZ x2490 = x2487 + x2489;
-Tbase TZ x2491 = 0x1a;
-Tbase TZ x2492 = x2490 >>> x2491;
-Tbase TZ x2493 = 0x1ffffff;
-Tbase TZ x2494 = x2253 & x2493;
-Tbase TZ x2495 = x2492 + x2494;
-Tbase TZ x2496 = 0x1ffffff;
-Tbase TZ x2497 = x2483 & x2496;
-Tbase TZ x2498 = 0x3ffffff;
-Tbase TZ x2499 = x2461 & x2498;
-Tbase TZ x2500 = 0x1ffffff;
-Tbase TZ x2501 = x2427 & x2500;
-Tbase TZ x2502 = 0x3ffffff;
-Tbase TZ x2503 = x2403 & x2502;
-Tbase TZ x2504 = 0x1ffffff;
-Tbase TZ x2505 = x2369 & x2504;
-Tbase TZ x2506 = 0x3ffffff;
-Tbase TZ x2507 = x2345 & x2506;
-Tbase TZ x2508 = 0x1ffffff;
-Tbase TZ x2509 = x2311 & x2508;
-Tbase TZ x2510 = 0x19;
-Tbase TZ x2511 = x2495 >>> x2510;
-Tbase TZ x2512 = 0x3ffffff;
-Tbase TZ x2513 = x2287 & x2512;
-Tbase TZ x2514 = x2511 + x2513;
-Tbase TZ x2515 = 0x1ffffff;
-Tbase TZ x2516 = x2495 & x2515;
-Tbase TZ x2517 = 0x3ffffff;
-Tbase TZ x2518 = x2490 & x2517;
-Tbase TZ x2519 = x418 * x768;
-Tbase TZ x2520 = 0x2;
-Tbase TZ x2521 = x766 * x2520;
-Tbase TZ x2522 = x397 * x2521;
-Tbase TZ x2523 = x399 * x764;
-Tbase TZ x2524 = 0x2;
-Tbase TZ x2525 = x759 * x2524;
-Tbase TZ x2526 = x401 * x2525;
-Tbase TZ x2527 = x403 * x757;
-Tbase TZ x2528 = 0x2;
-Tbase TZ x2529 = x755 * x2528;
-Tbase TZ x2530 = x405 * x2529;
-Tbase TZ x2531 = x407 * x753;
-Tbase TZ x2532 = 0x2;
-Tbase TZ x2533 = x751 * x2532;
-Tbase TZ x2534 = x409 * x2533;
-Tbase TZ x2535 = x414 * x749;
-Tbase TZ x2536 = 0x2;
-Tbase TZ x2537 = x747 * x2536;
-Tbase TZ x2538 = x416 * x2537;
-Tbase TZ x2539 = x2535 + x2538;
-Tbase TZ x2540 = x2534 + x2539;
-Tbase TZ x2541 = x2531 + x2540;
-Tbase TZ x2542 = x2530 + x2541;
-Tbase TZ x2543 = x2527 + x2542;
-Tbase TZ x2544 = x2526 + x2543;
-Tbase TZ x2545 = x2523 + x2544;
-Tbase TZ x2546 = x2522 + x2545;
-Tbase TZ x2547 = 0x13;
-Tbase TZ x2548 = x2547 * x2546;
-Tbase TZ x2549 = x2519 + x2548;
-Tbase TZ x2550 = 0x1a;
-Tbase TZ x2551 = x2549 >>> x2550;
-Tbase TZ x2552 = x416 * x768;
-Tbase TZ x2553 = x418 * x766;
-Tbase TZ x2554 = x2552 + x2553;
-Tbase TZ x2555 = x397 * x764;
-Tbase TZ x2556 = x399 * x759;
-Tbase TZ x2557 = x401 * x757;
-Tbase TZ x2558 = x403 * x755;
-Tbase TZ x2559 = x405 * x753;
-Tbase TZ x2560 = x407 * x751;
-Tbase TZ x2561 = x409 * x749;
-Tbase TZ x2562 = x414 * x747;
-Tbase TZ x2563 = x2561 + x2562;
-Tbase TZ x2564 = x2560 + x2563;
-Tbase TZ x2565 = x2559 + x2564;
-Tbase TZ x2566 = x2558 + x2565;
-Tbase TZ x2567 = x2557 + x2566;
-Tbase TZ x2568 = x2556 + x2567;
-Tbase TZ x2569 = x2555 + x2568;
-Tbase TZ x2570 = 0x13;
-Tbase TZ x2571 = x2570 * x2569;
-Tbase TZ x2572 = x2554 + x2571;
-Tbase TZ x2573 = x2551 + x2572;
-Tbase TZ x2574 = 0x19;
-Tbase TZ x2575 = x2573 >>> x2574;
-Tbase TZ x2576 = x414 * x768;
-Tbase TZ x2577 = 0x2;
-Tbase TZ x2578 = x766 * x2577;
-Tbase TZ x2579 = x416 * x2578;
-Tbase TZ x2580 = x418 * x764;
-Tbase TZ x2581 = x2579 + x2580;
-Tbase TZ x2582 = x2576 + x2581;
-Tbase TZ x2583 = 0x2;
-Tbase TZ x2584 = x759 * x2583;
-Tbase TZ x2585 = x397 * x2584;
-Tbase TZ x2586 = x399 * x757;
-Tbase TZ x2587 = 0x2;
-Tbase TZ x2588 = x755 * x2587;
-Tbase TZ x2589 = x401 * x2588;
-Tbase TZ x2590 = x403 * x753;
-Tbase TZ x2591 = 0x2;
-Tbase TZ x2592 = x751 * x2591;
-Tbase TZ x2593 = x405 * x2592;
-Tbase TZ x2594 = x407 * x749;
-Tbase TZ x2595 = 0x2;
-Tbase TZ x2596 = x747 * x2595;
-Tbase TZ x2597 = x409 * x2596;
-Tbase TZ x2598 = x2594 + x2597;
-Tbase TZ x2599 = x2593 + x2598;
-Tbase TZ x2600 = x2590 + x2599;
-Tbase TZ x2601 = x2589 + x2600;
-Tbase TZ x2602 = x2586 + x2601;
-Tbase TZ x2603 = x2585 + x2602;
-Tbase TZ x2604 = 0x13;
-Tbase TZ x2605 = x2604 * x2603;
-Tbase TZ x2606 = x2582 + x2605;
-Tbase TZ x2607 = x2575 + x2606;
-Tbase TZ x2608 = 0x1a;
-Tbase TZ x2609 = x2607 >>> x2608;
-Tbase TZ x2610 = x409 * x768;
-Tbase TZ x2611 = x414 * x766;
-Tbase TZ x2612 = x416 * x764;
-Tbase TZ x2613 = x418 * x759;
-Tbase TZ x2614 = x2612 + x2613;
-Tbase TZ x2615 = x2611 + x2614;
-Tbase TZ x2616 = x2610 + x2615;
-Tbase TZ x2617 = x397 * x757;
-Tbase TZ x2618 = x399 * x755;
-Tbase TZ x2619 = x401 * x753;
-Tbase TZ x2620 = x403 * x751;
-Tbase TZ x2621 = x405 * x749;
-Tbase TZ x2622 = x407 * x747;
-Tbase TZ x2623 = x2621 + x2622;
-Tbase TZ x2624 = x2620 + x2623;
-Tbase TZ x2625 = x2619 + x2624;
-Tbase TZ x2626 = x2618 + x2625;
-Tbase TZ x2627 = x2617 + x2626;
-Tbase TZ x2628 = 0x13;
-Tbase TZ x2629 = x2628 * x2627;
-Tbase TZ x2630 = x2616 + x2629;
-Tbase TZ x2631 = x2609 + x2630;
-Tbase TZ x2632 = 0x19;
-Tbase TZ x2633 = x2631 >>> x2632;
-Tbase TZ x2634 = x407 * x768;
-Tbase TZ x2635 = 0x2;
-Tbase TZ x2636 = x766 * x2635;
-Tbase TZ x2637 = x409 * x2636;
-Tbase TZ x2638 = x414 * x764;
-Tbase TZ x2639 = 0x2;
-Tbase TZ x2640 = x759 * x2639;
-Tbase TZ x2641 = x416 * x2640;
-Tbase TZ x2642 = x418 * x757;
-Tbase TZ x2643 = x2641 + x2642;
-Tbase TZ x2644 = x2638 + x2643;
-Tbase TZ x2645 = x2637 + x2644;
-Tbase TZ x2646 = x2634 + x2645;
-Tbase TZ x2647 = 0x2;
-Tbase TZ x2648 = x755 * x2647;
-Tbase TZ x2649 = x397 * x2648;
-Tbase TZ x2650 = x399 * x753;
-Tbase TZ x2651 = 0x2;
-Tbase TZ x2652 = x751 * x2651;
-Tbase TZ x2653 = x401 * x2652;
-Tbase TZ x2654 = x403 * x749;
-Tbase TZ x2655 = 0x2;
-Tbase TZ x2656 = x747 * x2655;
-Tbase TZ x2657 = x405 * x2656;
-Tbase TZ x2658 = x2654 + x2657;
-Tbase TZ x2659 = x2653 + x2658;
-Tbase TZ x2660 = x2650 + x2659;
-Tbase TZ x2661 = x2649 + x2660;
-Tbase TZ x2662 = 0x13;
-Tbase TZ x2663 = x2662 * x2661;
-Tbase TZ x2664 = x2646 + x2663;
-Tbase TZ x2665 = x2633 + x2664;
-Tbase TZ x2666 = 0x1a;
-Tbase TZ x2667 = x2665 >>> x2666;
-Tbase TZ x2668 = x405 * x768;
-Tbase TZ x2669 = x407 * x766;
-Tbase TZ x2670 = x409 * x764;
-Tbase TZ x2671 = x414 * x759;
-Tbase TZ x2672 = x416 * x757;
-Tbase TZ x2673 = x418 * x755;
-Tbase TZ x2674 = x2672 + x2673;
-Tbase TZ x2675 = x2671 + x2674;
-Tbase TZ x2676 = x2670 + x2675;
-Tbase TZ x2677 = x2669 + x2676;
-Tbase TZ x2678 = x2668 + x2677;
-Tbase TZ x2679 = x397 * x753;
-Tbase TZ x2680 = x399 * x751;
-Tbase TZ x2681 = x401 * x749;
-Tbase TZ x2682 = x403 * x747;
-Tbase TZ x2683 = x2681 + x2682;
-Tbase TZ x2684 = x2680 + x2683;
-Tbase TZ x2685 = x2679 + x2684;
-Tbase TZ x2686 = 0x13;
-Tbase TZ x2687 = x2686 * x2685;
-Tbase TZ x2688 = x2678 + x2687;
-Tbase TZ x2689 = x2667 + x2688;
-Tbase TZ x2690 = 0x19;
-Tbase TZ x2691 = x2689 >>> x2690;
-Tbase TZ x2692 = x403 * x768;
-Tbase TZ x2693 = 0x2;
-Tbase TZ x2694 = x766 * x2693;
-Tbase TZ x2695 = x405 * x2694;
-Tbase TZ x2696 = x407 * x764;
-Tbase TZ x2697 = 0x2;
-Tbase TZ x2698 = x759 * x2697;
-Tbase TZ x2699 = x409 * x2698;
-Tbase TZ x2700 = x414 * x757;
-Tbase TZ x2701 = 0x2;
-Tbase TZ x2702 = x755 * x2701;
-Tbase TZ x2703 = x416 * x2702;
-Tbase TZ x2704 = x418 * x753;
-Tbase TZ x2705 = x2703 + x2704;
-Tbase TZ x2706 = x2700 + x2705;
-Tbase TZ x2707 = x2699 + x2706;
-Tbase TZ x2708 = x2696 + x2707;
-Tbase TZ x2709 = x2695 + x2708;
-Tbase TZ x2710 = x2692 + x2709;
-Tbase TZ x2711 = 0x2;
-Tbase TZ x2712 = x751 * x2711;
-Tbase TZ x2713 = x397 * x2712;
-Tbase TZ x2714 = x399 * x749;
-Tbase TZ x2715 = 0x2;
-Tbase TZ x2716 = x747 * x2715;
-Tbase TZ x2717 = x401 * x2716;
-Tbase TZ x2718 = x2714 + x2717;
-Tbase TZ x2719 = x2713 + x2718;
-Tbase TZ x2720 = 0x13;
-Tbase TZ x2721 = x2720 * x2719;
-Tbase TZ x2722 = x2710 + x2721;
-Tbase TZ x2723 = x2691 + x2722;
-Tbase TZ x2724 = 0x1a;
-Tbase TZ x2725 = x2723 >>> x2724;
-Tbase TZ x2726 = x401 * x768;
-Tbase TZ x2727 = x403 * x766;
-Tbase TZ x2728 = x405 * x764;
-Tbase TZ x2729 = x407 * x759;
-Tbase TZ x2730 = x409 * x757;
-Tbase TZ x2731 = x414 * x755;
-Tbase TZ x2732 = x416 * x753;
-Tbase TZ x2733 = x418 * x751;
-Tbase TZ x2734 = x2732 + x2733;
-Tbase TZ x2735 = x2731 + x2734;
-Tbase TZ x2736 = x2730 + x2735;
-Tbase TZ x2737 = x2729 + x2736;
-Tbase TZ x2738 = x2728 + x2737;
-Tbase TZ x2739 = x2727 + x2738;
-Tbase TZ x2740 = x2726 + x2739;
-Tbase TZ x2741 = x397 * x749;
-Tbase TZ x2742 = x399 * x747;
-Tbase TZ x2743 = x2741 + x2742;
-Tbase TZ x2744 = 0x13;
-Tbase TZ x2745 = x2744 * x2743;
-Tbase TZ x2746 = x2740 + x2745;
-Tbase TZ x2747 = x2725 + x2746;
-Tbase TZ x2748 = 0x19;
-Tbase TZ x2749 = x2747 >>> x2748;
-Tbase TZ x2750 = x399 * x768;
-Tbase TZ x2751 = 0x2;
-Tbase TZ x2752 = x766 * x2751;
-Tbase TZ x2753 = x401 * x2752;
-Tbase TZ x2754 = x403 * x764;
-Tbase TZ x2755 = 0x2;
-Tbase TZ x2756 = x759 * x2755;
-Tbase TZ x2757 = x405 * x2756;
-Tbase TZ x2758 = x407 * x757;
-Tbase TZ x2759 = 0x2;
-Tbase TZ x2760 = x755 * x2759;
-Tbase TZ x2761 = x409 * x2760;
-Tbase TZ x2762 = x414 * x753;
-Tbase TZ x2763 = 0x2;
-Tbase TZ x2764 = x751 * x2763;
-Tbase TZ x2765 = x416 * x2764;
-Tbase TZ x2766 = x418 * x749;
-Tbase TZ x2767 = x2765 + x2766;
-Tbase TZ x2768 = x2762 + x2767;
-Tbase TZ x2769 = x2761 + x2768;
-Tbase TZ x2770 = x2758 + x2769;
-Tbase TZ x2771 = x2757 + x2770;
-Tbase TZ x2772 = x2754 + x2771;
-Tbase TZ x2773 = x2753 + x2772;
-Tbase TZ x2774 = x2750 + x2773;
-Tbase TZ x2775 = 0x2;
-Tbase TZ x2776 = x747 * x2775;
-Tbase TZ x2777 = x397 * x2776;
-Tbase TZ x2778 = 0x13;
-Tbase TZ x2779 = x2778 * x2777;
-Tbase TZ x2780 = x2774 + x2779;
-Tbase TZ x2781 = x2749 + x2780;
-Tbase TZ x2782 = 0x1a;
-Tbase TZ x2783 = x2781 >>> x2782;
-Tbase TZ x2784 = x397 * x768;
-Tbase TZ x2785 = x399 * x766;
-Tbase TZ x2786 = x401 * x764;
-Tbase TZ x2787 = x403 * x759;
-Tbase TZ x2788 = x405 * x757;
-Tbase TZ x2789 = x407 * x755;
-Tbase TZ x2790 = x409 * x753;
-Tbase TZ x2791 = x414 * x751;
-Tbase TZ x2792 = x416 * x749;
-Tbase TZ x2793 = x418 * x747;
-Tbase TZ x2794 = x2792 + x2793;
-Tbase TZ x2795 = x2791 + x2794;
-Tbase TZ x2796 = x2790 + x2795;
-Tbase TZ x2797 = x2789 + x2796;
-Tbase TZ x2798 = x2788 + x2797;
-Tbase TZ x2799 = x2787 + x2798;
-Tbase TZ x2800 = x2786 + x2799;
-Tbase TZ x2801 = x2785 + x2800;
-Tbase TZ x2802 = x2784 + x2801;
-Tbase TZ x2803 = x2783 + x2802;
-Tbase TZ x2804 = 0x19;
-Tbase TZ x2805 = x2803 >>> x2804;
-Tbase TZ x2806 = 0x13;
-Tbase TZ x2807 = x2806 * x2805;
-Tbase TZ x2808 = 0x3ffffff;
-Tbase TZ x2809 = x2549 & x2808;
-Tbase TZ x2810 = x2807 + x2809;
-Tbase TZ x2811 = 0x1a;
-Tbase TZ x2812 = x2810 >>> x2811;
-Tbase TZ x2813 = 0x1ffffff;
-Tbase TZ x2814 = x2573 & x2813;
-Tbase TZ x2815 = x2812 + x2814;
-Tbase TZ x2816 = 0x1ffffff;
-Tbase TZ x2817 = x2803 & x2816;
-Tbase TZ x2818 = 0x3ffffff;
-Tbase TZ x2819 = x2781 & x2818;
-Tbase TZ x2820 = 0x1ffffff;
-Tbase TZ x2821 = x2747 & x2820;
-Tbase TZ x2822 = 0x3ffffff;
-Tbase TZ x2823 = x2723 & x2822;
-Tbase TZ x2824 = 0x1ffffff;
-Tbase TZ x2825 = x2689 & x2824;
-Tbase TZ x2826 = 0x3ffffff;
-Tbase TZ x2827 = x2665 & x2826;
-Tbase TZ x2828 = 0x1ffffff;
-Tbase TZ x2829 = x2631 & x2828;
-Tbase TZ x2830 = 0x19;
-Tbase TZ x2831 = x2815 >>> x2830;
-Tbase TZ x2832 = 0x3ffffff;
-Tbase TZ x2833 = x2607 & x2832;
-Tbase TZ x2834 = x2831 + x2833;
-Tbase TZ x2835 = 0x1ffffff;
-Tbase TZ x2836 = x2815 & x2835;
-Tbase TZ x2837 = 0x3ffffff;
-Tbase TZ x2838 = x2810 & x2837;
-Tbase TZ x2839 = x38 * x798;
-Tbase TZ x2840 = 0x2;
-Tbase TZ x2841 = x795 * x2840;
-Tbase TZ x2842 = x29 * x2841;
-Tbase TZ x2843 = x30 * x792;
-Tbase TZ x2844 = 0x2;
-Tbase TZ x2845 = x789 * x2844;
-Tbase TZ x2846 = x31 * x2845;
-Tbase TZ x2847 = x32 * x786;
-Tbase TZ x2848 = 0x2;
-Tbase TZ x2849 = x783 * x2848;
-Tbase TZ x2850 = x33 * x2849;
-Tbase TZ x2851 = x34 * x780;
-Tbase TZ x2852 = 0x2;
-Tbase TZ x2853 = x777 * x2852;
-Tbase TZ x2854 = x35 * x2853;
-Tbase TZ x2855 = x36 * x774;
-Tbase TZ x2856 = 0x2;
-Tbase TZ x2857 = x771 * x2856;
-Tbase TZ x2858 = x37 * x2857;
-Tbase TZ x2859 = x2855 + x2858;
-Tbase TZ x2860 = x2854 + x2859;
-Tbase TZ x2861 = x2851 + x2860;
-Tbase TZ x2862 = x2850 + x2861;
-Tbase TZ x2863 = x2847 + x2862;
-Tbase TZ x2864 = x2846 + x2863;
-Tbase TZ x2865 = x2843 + x2864;
-Tbase TZ x2866 = x2842 + x2865;
-Tbase TZ x2867 = 0x13;
-Tbase TZ x2868 = x2867 * x2866;
-Tbase TZ x2869 = x2839 + x2868;
-Tbase TZ x2870 = 0x1a;
-Tbase TZ x2871 = x2869 >>> x2870;
-Tbase TZ x2872 = x37 * x798;
-Tbase TZ x2873 = x38 * x795;
-Tbase TZ x2874 = x2872 + x2873;
-Tbase TZ x2875 = x29 * x792;
-Tbase TZ x2876 = x30 * x789;
-Tbase TZ x2877 = x31 * x786;
-Tbase TZ x2878 = x32 * x783;
-Tbase TZ x2879 = x33 * x780;
-Tbase TZ x2880 = x34 * x777;
-Tbase TZ x2881 = x35 * x774;
-Tbase TZ x2882 = x36 * x771;
-Tbase TZ x2883 = x2881 + x2882;
-Tbase TZ x2884 = x2880 + x2883;
-Tbase TZ x2885 = x2879 + x2884;
-Tbase TZ x2886 = x2878 + x2885;
-Tbase TZ x2887 = x2877 + x2886;
-Tbase TZ x2888 = x2876 + x2887;
-Tbase TZ x2889 = x2875 + x2888;
-Tbase TZ x2890 = 0x13;
-Tbase TZ x2891 = x2890 * x2889;
-Tbase TZ x2892 = x2874 + x2891;
-Tbase TZ x2893 = x2871 + x2892;
-Tbase TZ x2894 = 0x19;
-Tbase TZ x2895 = x2893 >>> x2894;
-Tbase TZ x2896 = x36 * x798;
-Tbase TZ x2897 = 0x2;
-Tbase TZ x2898 = x795 * x2897;
-Tbase TZ x2899 = x37 * x2898;
-Tbase TZ x2900 = x38 * x792;
-Tbase TZ x2901 = x2899 + x2900;
-Tbase TZ x2902 = x2896 + x2901;
-Tbase TZ x2903 = 0x2;
-Tbase TZ x2904 = x789 * x2903;
-Tbase TZ x2905 = x29 * x2904;
-Tbase TZ x2906 = x30 * x786;
-Tbase TZ x2907 = 0x2;
-Tbase TZ x2908 = x783 * x2907;
-Tbase TZ x2909 = x31 * x2908;
-Tbase TZ x2910 = x32 * x780;
-Tbase TZ x2911 = 0x2;
-Tbase TZ x2912 = x777 * x2911;
-Tbase TZ x2913 = x33 * x2912;
-Tbase TZ x2914 = x34 * x774;
-Tbase TZ x2915 = 0x2;
-Tbase TZ x2916 = x771 * x2915;
-Tbase TZ x2917 = x35 * x2916;
-Tbase TZ x2918 = x2914 + x2917;
-Tbase TZ x2919 = x2913 + x2918;
-Tbase TZ x2920 = x2910 + x2919;
-Tbase TZ x2921 = x2909 + x2920;
-Tbase TZ x2922 = x2906 + x2921;
-Tbase TZ x2923 = x2905 + x2922;
-Tbase TZ x2924 = 0x13;
-Tbase TZ x2925 = x2924 * x2923;
-Tbase TZ x2926 = x2902 + x2925;
-Tbase TZ x2927 = x2895 + x2926;
-Tbase TZ x2928 = 0x1a;
-Tbase TZ x2929 = x2927 >>> x2928;
-Tbase TZ x2930 = x35 * x798;
-Tbase TZ x2931 = x36 * x795;
-Tbase TZ x2932 = x37 * x792;
-Tbase TZ x2933 = x38 * x789;
-Tbase TZ x2934 = x2932 + x2933;
-Tbase TZ x2935 = x2931 + x2934;
-Tbase TZ x2936 = x2930 + x2935;
-Tbase TZ x2937 = x29 * x786;
-Tbase TZ x2938 = x30 * x783;
-Tbase TZ x2939 = x31 * x780;
-Tbase TZ x2940 = x32 * x777;
-Tbase TZ x2941 = x33 * x774;
-Tbase TZ x2942 = x34 * x771;
-Tbase TZ x2943 = x2941 + x2942;
-Tbase TZ x2944 = x2940 + x2943;
-Tbase TZ x2945 = x2939 + x2944;
-Tbase TZ x2946 = x2938 + x2945;
-Tbase TZ x2947 = x2937 + x2946;
-Tbase TZ x2948 = 0x13;
-Tbase TZ x2949 = x2948 * x2947;
-Tbase TZ x2950 = x2936 + x2949;
-Tbase TZ x2951 = x2929 + x2950;
-Tbase TZ x2952 = 0x19;
-Tbase TZ x2953 = x2951 >>> x2952;
-Tbase TZ x2954 = x34 * x798;
-Tbase TZ x2955 = 0x2;
-Tbase TZ x2956 = x795 * x2955;
-Tbase TZ x2957 = x35 * x2956;
-Tbase TZ x2958 = x36 * x792;
-Tbase TZ x2959 = 0x2;
-Tbase TZ x2960 = x789 * x2959;
-Tbase TZ x2961 = x37 * x2960;
-Tbase TZ x2962 = x38 * x786;
-Tbase TZ x2963 = x2961 + x2962;
-Tbase TZ x2964 = x2958 + x2963;
-Tbase TZ x2965 = x2957 + x2964;
-Tbase TZ x2966 = x2954 + x2965;
-Tbase TZ x2967 = 0x2;
-Tbase TZ x2968 = x783 * x2967;
-Tbase TZ x2969 = x29 * x2968;
-Tbase TZ x2970 = x30 * x780;
-Tbase TZ x2971 = 0x2;
-Tbase TZ x2972 = x777 * x2971;
-Tbase TZ x2973 = x31 * x2972;
-Tbase TZ x2974 = x32 * x774;
-Tbase TZ x2975 = 0x2;
-Tbase TZ x2976 = x771 * x2975;
-Tbase TZ x2977 = x33 * x2976;
-Tbase TZ x2978 = x2974 + x2977;
-Tbase TZ x2979 = x2973 + x2978;
-Tbase TZ x2980 = x2970 + x2979;
-Tbase TZ x2981 = x2969 + x2980;
-Tbase TZ x2982 = 0x13;
-Tbase TZ x2983 = x2982 * x2981;
-Tbase TZ x2984 = x2966 + x2983;
-Tbase TZ x2985 = x2953 + x2984;
-Tbase TZ x2986 = 0x1a;
-Tbase TZ x2987 = x2985 >>> x2986;
-Tbase TZ x2988 = x33 * x798;
-Tbase TZ x2989 = x34 * x795;
-Tbase TZ x2990 = x35 * x792;
-Tbase TZ x2991 = x36 * x789;
-Tbase TZ x2992 = x37 * x786;
-Tbase TZ x2993 = x38 * x783;
-Tbase TZ x2994 = x2992 + x2993;
-Tbase TZ x2995 = x2991 + x2994;
-Tbase TZ x2996 = x2990 + x2995;
-Tbase TZ x2997 = x2989 + x2996;
-Tbase TZ x2998 = x2988 + x2997;
-Tbase TZ x2999 = x29 * x780;
-Tbase TZ x3000 = x30 * x777;
-Tbase TZ x3001 = x31 * x774;
-Tbase TZ x3002 = x32 * x771;
-Tbase TZ x3003 = x3001 + x3002;
-Tbase TZ x3004 = x3000 + x3003;
-Tbase TZ x3005 = x2999 + x3004;
-Tbase TZ x3006 = 0x13;
-Tbase TZ x3007 = x3006 * x3005;
-Tbase TZ x3008 = x2998 + x3007;
-Tbase TZ x3009 = x2987 + x3008;
-Tbase TZ x3010 = 0x19;
-Tbase TZ x3011 = x3009 >>> x3010;
-Tbase TZ x3012 = x32 * x798;
-Tbase TZ x3013 = 0x2;
-Tbase TZ x3014 = x795 * x3013;
-Tbase TZ x3015 = x33 * x3014;
-Tbase TZ x3016 = x34 * x792;
-Tbase TZ x3017 = 0x2;
-Tbase TZ x3018 = x789 * x3017;
-Tbase TZ x3019 = x35 * x3018;
-Tbase TZ x3020 = x36 * x786;
-Tbase TZ x3021 = 0x2;
-Tbase TZ x3022 = x783 * x3021;
-Tbase TZ x3023 = x37 * x3022;
-Tbase TZ x3024 = x38 * x780;
-Tbase TZ x3025 = x3023 + x3024;
-Tbase TZ x3026 = x3020 + x3025;
-Tbase TZ x3027 = x3019 + x3026;
-Tbase TZ x3028 = x3016 + x3027;
-Tbase TZ x3029 = x3015 + x3028;
-Tbase TZ x3030 = x3012 + x3029;
-Tbase TZ x3031 = 0x2;
-Tbase TZ x3032 = x777 * x3031;
-Tbase TZ x3033 = x29 * x3032;
-Tbase TZ x3034 = x30 * x774;
-Tbase TZ x3035 = 0x2;
-Tbase TZ x3036 = x771 * x3035;
-Tbase TZ x3037 = x31 * x3036;
-Tbase TZ x3038 = x3034 + x3037;
-Tbase TZ x3039 = x3033 + x3038;
-Tbase TZ x3040 = 0x13;
-Tbase TZ x3041 = x3040 * x3039;
-Tbase TZ x3042 = x3030 + x3041;
-Tbase TZ x3043 = x3011 + x3042;
-Tbase TZ x3044 = 0x1a;
-Tbase TZ x3045 = x3043 >>> x3044;
-Tbase TZ x3046 = x31 * x798;
-Tbase TZ x3047 = x32 * x795;
-Tbase TZ x3048 = x33 * x792;
-Tbase TZ x3049 = x34 * x789;
-Tbase TZ x3050 = x35 * x786;
-Tbase TZ x3051 = x36 * x783;
-Tbase TZ x3052 = x37 * x780;
-Tbase TZ x3053 = x38 * x777;
-Tbase TZ x3054 = x3052 + x3053;
-Tbase TZ x3055 = x3051 + x3054;
-Tbase TZ x3056 = x3050 + x3055;
-Tbase TZ x3057 = x3049 + x3056;
-Tbase TZ x3058 = x3048 + x3057;
-Tbase TZ x3059 = x3047 + x3058;
-Tbase TZ x3060 = x3046 + x3059;
-Tbase TZ x3061 = x29 * x774;
-Tbase TZ x3062 = x30 * x771;
-Tbase TZ x3063 = x3061 + x3062;
-Tbase TZ x3064 = 0x13;
-Tbase TZ x3065 = x3064 * x3063;
-Tbase TZ x3066 = x3060 + x3065;
-Tbase TZ x3067 = x3045 + x3066;
-Tbase TZ x3068 = 0x19;
-Tbase TZ x3069 = x3067 >>> x3068;
-Tbase TZ x3070 = x30 * x798;
-Tbase TZ x3071 = 0x2;
-Tbase TZ x3072 = x795 * x3071;
-Tbase TZ x3073 = x31 * x3072;
-Tbase TZ x3074 = x32 * x792;
-Tbase TZ x3075 = 0x2;
-Tbase TZ x3076 = x789 * x3075;
-Tbase TZ x3077 = x33 * x3076;
-Tbase TZ x3078 = x34 * x786;
-Tbase TZ x3079 = 0x2;
-Tbase TZ x3080 = x783 * x3079;
-Tbase TZ x3081 = x35 * x3080;
-Tbase TZ x3082 = x36 * x780;
-Tbase TZ x3083 = 0x2;
-Tbase TZ x3084 = x777 * x3083;
-Tbase TZ x3085 = x37 * x3084;
-Tbase TZ x3086 = x38 * x774;
-Tbase TZ x3087 = x3085 + x3086;
-Tbase TZ x3088 = x3082 + x3087;
-Tbase TZ x3089 = x3081 + x3088;
-Tbase TZ x3090 = x3078 + x3089;
-Tbase TZ x3091 = x3077 + x3090;
-Tbase TZ x3092 = x3074 + x3091;
-Tbase TZ x3093 = x3073 + x3092;
-Tbase TZ x3094 = x3070 + x3093;
-Tbase TZ x3095 = 0x2;
-Tbase TZ x3096 = x771 * x3095;
-Tbase TZ x3097 = x29 * x3096;
-Tbase TZ x3098 = 0x13;
-Tbase TZ x3099 = x3098 * x3097;
-Tbase TZ x3100 = x3094 + x3099;
-Tbase TZ x3101 = x3069 + x3100;
-Tbase TZ x3102 = 0x1a;
-Tbase TZ x3103 = x3101 >>> x3102;
-Tbase TZ x3104 = x29 * x798;
-Tbase TZ x3105 = x30 * x795;
-Tbase TZ x3106 = x31 * x792;
-Tbase TZ x3107 = x32 * x789;
-Tbase TZ x3108 = x33 * x786;
-Tbase TZ x3109 = x34 * x783;
-Tbase TZ x3110 = x35 * x780;
-Tbase TZ x3111 = x36 * x777;
-Tbase TZ x3112 = x37 * x774;
-Tbase TZ x3113 = x38 * x771;
-Tbase TZ x3114 = x3112 + x3113;
-Tbase TZ x3115 = x3111 + x3114;
-Tbase TZ x3116 = x3110 + x3115;
-Tbase TZ x3117 = x3109 + x3116;
-Tbase TZ x3118 = x3108 + x3117;
-Tbase TZ x3119 = x3107 + x3118;
-Tbase TZ x3120 = x3106 + x3119;
-Tbase TZ x3121 = x3105 + x3120;
-Tbase TZ x3122 = x3104 + x3121;
-Tbase TZ x3123 = x3103 + x3122;
-Tbase TZ x3124 = 0x19;
-Tbase TZ x3125 = x3123 >>> x3124;
-Tbase TZ x3126 = 0x13;
-Tbase TZ x3127 = x3126 * x3125;
-Tbase TZ x3128 = 0x3ffffff;
-Tbase TZ x3129 = x2869 & x3128;
-Tbase TZ x3130 = x3127 + x3129;
-Tbase TZ x3131 = 0x1a;
-Tbase TZ x3132 = x3130 >>> x3131;
-Tbase TZ x3133 = 0x1ffffff;
-Tbase TZ x3134 = x2893 & x3133;
-Tbase TZ x3135 = x3132 + x3134;
-Tbase TZ x3136 = 0x1ffffff;
-Tbase TZ x3137 = x3123 & x3136;
-Tbase TZ x3138 = 0x3ffffff;
-Tbase TZ x3139 = x3101 & x3138;
-Tbase TZ x3140 = 0x1ffffff;
-Tbase TZ x3141 = x3067 & x3140;
-Tbase TZ x3142 = 0x3ffffff;
-Tbase TZ x3143 = x3043 & x3142;
-Tbase TZ x3144 = 0x1ffffff;
-Tbase TZ x3145 = x3009 & x3144;
-Tbase TZ x3146 = 0x3ffffff;
-Tbase TZ x3147 = x2985 & x3146;
-Tbase TZ x3148 = 0x1ffffff;
-Tbase TZ x3149 = x2951 & x3148;
-Tbase TZ x3150 = 0x19;
-Tbase TZ x3151 = x3135 >>> x3150;
-Tbase TZ x3152 = 0x3ffffff;
-Tbase TZ x3153 = x2927 & x3152;
-Tbase TZ x3154 = x3151 + x3153;
-Tbase TZ x3155 = 0x1ffffff;
-Tbase TZ x3156 = x3135 & x3155;
-Tbase TZ x3157 = 0x3ffffff;
-Tbase TZ x3158 = x3130 & x3157;
-Tbase TZ x3159 = x397 + x3137;
-Tbase TZ x3160 = x399 + x3139;
-Tbase TZ x3161 = x401 + x3141;
-Tbase TZ x3162 = x403 + x3143;
-Tbase TZ x3163 = x405 + x3145;
-Tbase TZ x3164 = x407 + x3147;
-Tbase TZ x3165 = x409 + x3149;
-Tbase TZ x3166 = x414 + x3154;
-Tbase TZ x3167 = x416 + x3156;
-Tbase TZ x3168 = x418 + x3158;
-Tbase TZ x3169 = x798 * x3168;
-Tbase TZ x3170 = 0x2;
-Tbase TZ x3171 = x3167 * x3170;
-Tbase TZ x3172 = x771 * x3171;
-Tbase TZ x3173 = x774 * x3166;
-Tbase TZ x3174 = 0x2;
-Tbase TZ x3175 = x3165 * x3174;
-Tbase TZ x3176 = x777 * x3175;
-Tbase TZ x3177 = x780 * x3164;
-Tbase TZ x3178 = 0x2;
-Tbase TZ x3179 = x3163 * x3178;
-Tbase TZ x3180 = x783 * x3179;
-Tbase TZ x3181 = x786 * x3162;
-Tbase TZ x3182 = 0x2;
-Tbase TZ x3183 = x3161 * x3182;
-Tbase TZ x3184 = x789 * x3183;
-Tbase TZ x3185 = x792 * x3160;
-Tbase TZ x3186 = 0x2;
-Tbase TZ x3187 = x3159 * x3186;
-Tbase TZ x3188 = x795 * x3187;
-Tbase TZ x3189 = x3185 + x3188;
-Tbase TZ x3190 = x3184 + x3189;
-Tbase TZ x3191 = x3181 + x3190;
-Tbase TZ x3192 = x3180 + x3191;
-Tbase TZ x3193 = x3177 + x3192;
-Tbase TZ x3194 = x3176 + x3193;
-Tbase TZ x3195 = x3173 + x3194;
-Tbase TZ x3196 = x3172 + x3195;
-Tbase TZ x3197 = 0x13;
-Tbase TZ x3198 = x3197 * x3196;
-Tbase TZ x3199 = x3169 + x3198;
-Tbase TZ x3200 = 0x1a;
-Tbase TZ x3201 = x3199 >>> x3200;
-Tbase TZ x3202 = x795 * x3168;
-Tbase TZ x3203 = x798 * x3167;
-Tbase TZ x3204 = x3202 + x3203;
-Tbase TZ x3205 = x771 * x3166;
-Tbase TZ x3206 = x774 * x3165;
-Tbase TZ x3207 = x777 * x3164;
-Tbase TZ x3208 = x780 * x3163;
-Tbase TZ x3209 = x783 * x3162;
-Tbase TZ x3210 = x786 * x3161;
-Tbase TZ x3211 = x789 * x3160;
-Tbase TZ x3212 = x792 * x3159;
-Tbase TZ x3213 = x3211 + x3212;
-Tbase TZ x3214 = x3210 + x3213;
-Tbase TZ x3215 = x3209 + x3214;
-Tbase TZ x3216 = x3208 + x3215;
-Tbase TZ x3217 = x3207 + x3216;
-Tbase TZ x3218 = x3206 + x3217;
-Tbase TZ x3219 = x3205 + x3218;
-Tbase TZ x3220 = 0x13;
-Tbase TZ x3221 = x3220 * x3219;
-Tbase TZ x3222 = x3204 + x3221;
-Tbase TZ x3223 = x3201 + x3222;
-Tbase TZ x3224 = 0x19;
-Tbase TZ x3225 = x3223 >>> x3224;
-Tbase TZ x3226 = x792 * x3168;
-Tbase TZ x3227 = 0x2;
-Tbase TZ x3228 = x3167 * x3227;
-Tbase TZ x3229 = x795 * x3228;
-Tbase TZ x3230 = x798 * x3166;
-Tbase TZ x3231 = x3229 + x3230;
-Tbase TZ x3232 = x3226 + x3231;
-Tbase TZ x3233 = 0x2;
-Tbase TZ x3234 = x3165 * x3233;
-Tbase TZ x3235 = x771 * x3234;
-Tbase TZ x3236 = x774 * x3164;
-Tbase TZ x3237 = 0x2;
-Tbase TZ x3238 = x3163 * x3237;
-Tbase TZ x3239 = x777 * x3238;
-Tbase TZ x3240 = x780 * x3162;
-Tbase TZ x3241 = 0x2;
-Tbase TZ x3242 = x3161 * x3241;
-Tbase TZ x3243 = x783 * x3242;
-Tbase TZ x3244 = x786 * x3160;
-Tbase TZ x3245 = 0x2;
-Tbase TZ x3246 = x3159 * x3245;
-Tbase TZ x3247 = x789 * x3246;
-Tbase TZ x3248 = x3244 + x3247;
-Tbase TZ x3249 = x3243 + x3248;
-Tbase TZ x3250 = x3240 + x3249;
-Tbase TZ x3251 = x3239 + x3250;
-Tbase TZ x3252 = x3236 + x3251;
-Tbase TZ x3253 = x3235 + x3252;
-Tbase TZ x3254 = 0x13;
-Tbase TZ x3255 = x3254 * x3253;
-Tbase TZ x3256 = x3232 + x3255;
-Tbase TZ x3257 = x3225 + x3256;
-Tbase TZ x3258 = 0x1a;
-Tbase TZ x3259 = x3257 >>> x3258;
-Tbase TZ x3260 = x789 * x3168;
-Tbase TZ x3261 = x792 * x3167;
-Tbase TZ x3262 = x795 * x3166;
-Tbase TZ x3263 = x798 * x3165;
-Tbase TZ x3264 = x3262 + x3263;
-Tbase TZ x3265 = x3261 + x3264;
-Tbase TZ x3266 = x3260 + x3265;
-Tbase TZ x3267 = x771 * x3164;
-Tbase TZ x3268 = x774 * x3163;
-Tbase TZ x3269 = x777 * x3162;
-Tbase TZ x3270 = x780 * x3161;
-Tbase TZ x3271 = x783 * x3160;
-Tbase TZ x3272 = x786 * x3159;
-Tbase TZ x3273 = x3271 + x3272;
-Tbase TZ x3274 = x3270 + x3273;
-Tbase TZ x3275 = x3269 + x3274;
-Tbase TZ x3276 = x3268 + x3275;
-Tbase TZ x3277 = x3267 + x3276;
-Tbase TZ x3278 = 0x13;
-Tbase TZ x3279 = x3278 * x3277;
-Tbase TZ x3280 = x3266 + x3279;
-Tbase TZ x3281 = x3259 + x3280;
-Tbase TZ x3282 = 0x19;
-Tbase TZ x3283 = x3281 >>> x3282;
-Tbase TZ x3284 = x786 * x3168;
-Tbase TZ x3285 = 0x2;
-Tbase TZ x3286 = x3167 * x3285;
-Tbase TZ x3287 = x789 * x3286;
-Tbase TZ x3288 = x792 * x3166;
-Tbase TZ x3289 = 0x2;
-Tbase TZ x3290 = x3165 * x3289;
-Tbase TZ x3291 = x795 * x3290;
-Tbase TZ x3292 = x798 * x3164;
-Tbase TZ x3293 = x3291 + x3292;
-Tbase TZ x3294 = x3288 + x3293;
-Tbase TZ x3295 = x3287 + x3294;
-Tbase TZ x3296 = x3284 + x3295;
-Tbase TZ x3297 = 0x2;
-Tbase TZ x3298 = x3163 * x3297;
-Tbase TZ x3299 = x771 * x3298;
-Tbase TZ x3300 = x774 * x3162;
-Tbase TZ x3301 = 0x2;
-Tbase TZ x3302 = x3161 * x3301;
-Tbase TZ x3303 = x777 * x3302;
-Tbase TZ x3304 = x780 * x3160;
-Tbase TZ x3305 = 0x2;
-Tbase TZ x3306 = x3159 * x3305;
-Tbase TZ x3307 = x783 * x3306;
-Tbase TZ x3308 = x3304 + x3307;
-Tbase TZ x3309 = x3303 + x3308;
-Tbase TZ x3310 = x3300 + x3309;
-Tbase TZ x3311 = x3299 + x3310;
-Tbase TZ x3312 = 0x13;
-Tbase TZ x3313 = x3312 * x3311;
-Tbase TZ x3314 = x3296 + x3313;
-Tbase TZ x3315 = x3283 + x3314;
-Tbase TZ x3316 = 0x1a;
-Tbase TZ x3317 = x3315 >>> x3316;
-Tbase TZ x3318 = x783 * x3168;
-Tbase TZ x3319 = x786 * x3167;
-Tbase TZ x3320 = x789 * x3166;
-Tbase TZ x3321 = x792 * x3165;
-Tbase TZ x3322 = x795 * x3164;
-Tbase TZ x3323 = x798 * x3163;
-Tbase TZ x3324 = x3322 + x3323;
-Tbase TZ x3325 = x3321 + x3324;
-Tbase TZ x3326 = x3320 + x3325;
-Tbase TZ x3327 = x3319 + x3326;
-Tbase TZ x3328 = x3318 + x3327;
-Tbase TZ x3329 = x771 * x3162;
-Tbase TZ x3330 = x774 * x3161;
-Tbase TZ x3331 = x777 * x3160;
-Tbase TZ x3332 = x780 * x3159;
-Tbase TZ x3333 = x3331 + x3332;
-Tbase TZ x3334 = x3330 + x3333;
-Tbase TZ x3335 = x3329 + x3334;
-Tbase TZ x3336 = 0x13;
-Tbase TZ x3337 = x3336 * x3335;
-Tbase TZ x3338 = x3328 + x3337;
-Tbase TZ x3339 = x3317 + x3338;
-Tbase TZ x3340 = 0x19;
-Tbase TZ x3341 = x3339 >>> x3340;
-Tbase TZ x3342 = x780 * x3168;
-Tbase TZ x3343 = 0x2;
-Tbase TZ x3344 = x3167 * x3343;
-Tbase TZ x3345 = x783 * x3344;
-Tbase TZ x3346 = x786 * x3166;
-Tbase TZ x3347 = 0x2;
-Tbase TZ x3348 = x3165 * x3347;
-Tbase TZ x3349 = x789 * x3348;
-Tbase TZ x3350 = x792 * x3164;
-Tbase TZ x3351 = 0x2;
-Tbase TZ x3352 = x3163 * x3351;
-Tbase TZ x3353 = x795 * x3352;
-Tbase TZ x3354 = x798 * x3162;
-Tbase TZ x3355 = x3353 + x3354;
-Tbase TZ x3356 = x3350 + x3355;
-Tbase TZ x3357 = x3349 + x3356;
-Tbase TZ x3358 = x3346 + x3357;
-Tbase TZ x3359 = x3345 + x3358;
-Tbase TZ x3360 = x3342 + x3359;
-Tbase TZ x3361 = 0x2;
-Tbase TZ x3362 = x3161 * x3361;
-Tbase TZ x3363 = x771 * x3362;
-Tbase TZ x3364 = x774 * x3160;
-Tbase TZ x3365 = 0x2;
-Tbase TZ x3366 = x3159 * x3365;
-Tbase TZ x3367 = x777 * x3366;
-Tbase TZ x3368 = x3364 + x3367;
-Tbase TZ x3369 = x3363 + x3368;
-Tbase TZ x3370 = 0x13;
-Tbase TZ x3371 = x3370 * x3369;
-Tbase TZ x3372 = x3360 + x3371;
-Tbase TZ x3373 = x3341 + x3372;
-Tbase TZ x3374 = 0x1a;
-Tbase TZ x3375 = x3373 >>> x3374;
-Tbase TZ x3376 = x777 * x3168;
-Tbase TZ x3377 = x780 * x3167;
-Tbase TZ x3378 = x783 * x3166;
-Tbase TZ x3379 = x786 * x3165;
-Tbase TZ x3380 = x789 * x3164;
-Tbase TZ x3381 = x792 * x3163;
-Tbase TZ x3382 = x795 * x3162;
-Tbase TZ x3383 = x798 * x3161;
-Tbase TZ x3384 = x3382 + x3383;
-Tbase TZ x3385 = x3381 + x3384;
-Tbase TZ x3386 = x3380 + x3385;
-Tbase TZ x3387 = x3379 + x3386;
-Tbase TZ x3388 = x3378 + x3387;
-Tbase TZ x3389 = x3377 + x3388;
-Tbase TZ x3390 = x3376 + x3389;
-Tbase TZ x3391 = x771 * x3160;
-Tbase TZ x3392 = x774 * x3159;
-Tbase TZ x3393 = x3391 + x3392;
-Tbase TZ x3394 = 0x13;
-Tbase TZ x3395 = x3394 * x3393;
-Tbase TZ x3396 = x3390 + x3395;
-Tbase TZ x3397 = x3375 + x3396;
-Tbase TZ x3398 = 0x19;
-Tbase TZ x3399 = x3397 >>> x3398;
-Tbase TZ x3400 = x774 * x3168;
-Tbase TZ x3401 = 0x2;
-Tbase TZ x3402 = x3167 * x3401;
-Tbase TZ x3403 = x777 * x3402;
-Tbase TZ x3404 = x780 * x3166;
-Tbase TZ x3405 = 0x2;
-Tbase TZ x3406 = x3165 * x3405;
-Tbase TZ x3407 = x783 * x3406;
-Tbase TZ x3408 = x786 * x3164;
-Tbase TZ x3409 = 0x2;
-Tbase TZ x3410 = x3163 * x3409;
-Tbase TZ x3411 = x789 * x3410;
-Tbase TZ x3412 = x792 * x3162;
-Tbase TZ x3413 = 0x2;
-Tbase TZ x3414 = x3161 * x3413;
-Tbase TZ x3415 = x795 * x3414;
-Tbase TZ x3416 = x798 * x3160;
-Tbase TZ x3417 = x3415 + x3416;
-Tbase TZ x3418 = x3412 + x3417;
-Tbase TZ x3419 = x3411 + x3418;
-Tbase TZ x3420 = x3408 + x3419;
-Tbase TZ x3421 = x3407 + x3420;
-Tbase TZ x3422 = x3404 + x3421;
-Tbase TZ x3423 = x3403 + x3422;
-Tbase TZ x3424 = x3400 + x3423;
-Tbase TZ x3425 = 0x2;
-Tbase TZ x3426 = x3159 * x3425;
-Tbase TZ x3427 = x771 * x3426;
-Tbase TZ x3428 = 0x13;
-Tbase TZ x3429 = x3428 * x3427;
-Tbase TZ x3430 = x3424 + x3429;
-Tbase TZ x3431 = x3399 + x3430;
-Tbase TZ x3432 = 0x1a;
-Tbase TZ x3433 = x3431 >>> x3432;
-Tbase TZ x3434 = x771 * x3168;
-Tbase TZ x3435 = x774 * x3167;
-Tbase TZ x3436 = x777 * x3166;
-Tbase TZ x3437 = x780 * x3165;
-Tbase TZ x3438 = x783 * x3164;
-Tbase TZ x3439 = x786 * x3163;
-Tbase TZ x3440 = x789 * x3162;
-Tbase TZ x3441 = x792 * x3161;
-Tbase TZ x3442 = x795 * x3160;
-Tbase TZ x3443 = x798 * x3159;
-Tbase TZ x3444 = x3442 + x3443;
-Tbase TZ x3445 = x3441 + x3444;
-Tbase TZ x3446 = x3440 + x3445;
-Tbase TZ x3447 = x3439 + x3446;
-Tbase TZ x3448 = x3438 + x3447;
-Tbase TZ x3449 = x3437 + x3448;
-Tbase TZ x3450 = x3436 + x3449;
-Tbase TZ x3451 = x3435 + x3450;
-Tbase TZ x3452 = x3434 + x3451;
-Tbase TZ x3453 = x3433 + x3452;
-Tbase TZ x3454 = 0x19;
-Tbase TZ x3455 = x3453 >>> x3454;
-Tbase TZ x3456 = 0x13;
-Tbase TZ x3457 = x3456 * x3455;
-Tbase TZ x3458 = 0x3ffffff;
-Tbase TZ x3459 = x3199 & x3458;
-Tbase TZ x3460 = x3457 + x3459;
-Tbase TZ x3461 = 0x1a;
-Tbase TZ x3462 = x3460 >>> x3461;
-Tbase TZ x3463 = 0x1ffffff;
-Tbase TZ x3464 = x3223 & x3463;
-Tbase TZ x3465 = x3462 + x3464;
-Tbase TZ x3466 = 0x1ffffff;
-Tbase TZ x3467 = x3453 & x3466;
-Tbase TZ x3468 = 0x3ffffff;
-Tbase TZ x3469 = x3431 & x3468;
-Tbase TZ x3470 = 0x1ffffff;
-Tbase TZ x3471 = x3397 & x3470;
-Tbase TZ x3472 = 0x3ffffff;
-Tbase TZ x3473 = x3373 & x3472;
-Tbase TZ x3474 = 0x1ffffff;
-Tbase TZ x3475 = x3339 & x3474;
-Tbase TZ x3476 = 0x3ffffff;
-Tbase TZ x3477 = x3315 & x3476;
-Tbase TZ x3478 = 0x1ffffff;
-Tbase TZ x3479 = x3281 & x3478;
-Tbase TZ x3480 = 0x19;
-Tbase TZ x3481 = x3465 >>> x3480;
-Tbase TZ x3482 = 0x3ffffff;
-Tbase TZ x3483 = x3257 & x3482;
-Tbase TZ x3484 = x3481 + x3483;
-Tbase TZ x3485 = 0x1ffffff;
-Tbase TZ x3486 = x3465 & x3485;
-Tbase TZ x3487 = 0x3ffffff;
-Tbase TZ x3488 = x3460 & x3487;
-(Return x2817, Return x2819, Return x2821, Return x2823,
-Return x2825, Return x2827, Return x2829, Return x2834,
-Return x2836, Return x2838,
-(Return x3467, Return x3469, Return x3471, Return x3473,
-Return x3475, Return x3477, Return x3479, Return x3484,
-Return x3486, Return x3488),
-(Return x1797, Return x1799, Return x1801, Return x1803,
-Return x1805, Return x1807, Return x1809, Return x1814,
-Return x1816, Return x1818),
-(Return x2497, Return x2499, Return x2501, Return x2503,
-Return x2505, Return x2507, Return x2509, Return x2514,
-Return x2516, Return x2518))
- : forall var : base_type -> Type,
- expr base_type op
- (TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ ->
- TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> TZ -> ...)
-
-Argument scope is [function_scope]
diff --git a/src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.v b/src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.v
deleted file mode 100644
index 34afd08b6..000000000
--- a/src/Specific/GF25519Reflective/Reified/LadderStepJavaDisplay.v
+++ /dev/null
@@ -1,4 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.Reified.LadderStep.
-Require Export Crypto.Reflection.Z.JavaNotations.
-
-Print rladderstepW.
diff --git a/src/Specific/GF25519Reflective/Reified/Mul.v b/src/Specific/GF25519Reflective/Reified/Mul.v
deleted file mode 100644
index 0388c626a..000000000
--- a/src/Specific/GF25519Reflective/Reified/Mul.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonBinOp.
-
-Definition rmulZ_sig : rexpr_binop_sig mul. Proof. reify_sig. Defined.
-Definition rmulW := Eval vm_compute in rword_of_Z rmulZ_sig.
-Lemma rmulW_correct_and_bounded_gen : correct_and_bounded_genT rmulW rmulZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rmul_output_bounds := Eval vm_compute in compute_bounds rmulW ExprBinOp_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition rmulW_correct_and_bounded
- := ExprBinOp_correct_and_bounded
- rmulW mul rmulZ_sig rmulW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("Mul", compute_bounds_for_display rmulW ExprBinOp_bounds).
-Compute ("Mul overflows? ", sanity_compute rmulW ExprBinOp_bounds).
-Compute ("Mul overflows (error if it does)? ", sanity_check rmulW ExprBinOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/MulDisplay.log b/src/Specific/GF25519Reflective/Reified/MulDisplay.log
deleted file mode 100644
index 30f70a7a2..000000000
--- a/src/Specific/GF25519Reflective/Reified/MulDisplay.log
+++ /dev/null
@@ -1,297 +0,0 @@
-rmulW =
-fun var : Syntax.base_type -> Type =>
-x x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17
- x18 : var Syntax.TZ,
-Tbase Syntax.TZ x19 = x8 * x18;
-Tbase Syntax.TZ x20 = x17 * 0x2;
-Tbase Syntax.TZ x21 = x * x20;
-Tbase Syntax.TZ x22 = x0 * x16;
-Tbase Syntax.TZ x23 = x15 * 0x2;
-Tbase Syntax.TZ x24 = x1 * x23;
-Tbase Syntax.TZ x25 = x2 * x14;
-Tbase Syntax.TZ x26 = x13 * 0x2;
-Tbase Syntax.TZ x27 = x3 * x26;
-Tbase Syntax.TZ x28 = x4 * x12;
-Tbase Syntax.TZ x29 = x11 * 0x2;
-Tbase Syntax.TZ x30 = x5 * x29;
-Tbase Syntax.TZ x31 = x6 * x10;
-Tbase Syntax.TZ x32 = x9 * 0x2;
-Tbase Syntax.TZ x33 = x7 * x32;
-Tbase Syntax.TZ x34 = x31 + x33;
-Tbase Syntax.TZ x35 = x30 + x34;
-Tbase Syntax.TZ x36 = x28 + x35;
-Tbase Syntax.TZ x37 = x27 + x36;
-Tbase Syntax.TZ x38 = x25 + x37;
-Tbase Syntax.TZ x39 = x24 + x38;
-Tbase Syntax.TZ x40 = x22 + x39;
-Tbase Syntax.TZ x41 = x21 + x40;
-Tbase Syntax.TZ x42 = 0x13 * x41;
-Tbase Syntax.TZ x43 = x19 + x42;
-Tbase Syntax.TZ x44 = x43 >> 0x1a;
-Tbase Syntax.TZ x45 = x7 * x18;
-Tbase Syntax.TZ x46 = x8 * x17;
-Tbase Syntax.TZ x47 = x45 + x46;
-Tbase Syntax.TZ x48 = x * x16;
-Tbase Syntax.TZ x49 = x0 * x15;
-Tbase Syntax.TZ x50 = x1 * x14;
-Tbase Syntax.TZ x51 = x2 * x13;
-Tbase Syntax.TZ x52 = x3 * x12;
-Tbase Syntax.TZ x53 = x4 * x11;
-Tbase Syntax.TZ x54 = x5 * x10;
-Tbase Syntax.TZ x55 = x6 * x9;
-Tbase Syntax.TZ x56 = x54 + x55;
-Tbase Syntax.TZ x57 = x53 + x56;
-Tbase Syntax.TZ x58 = x52 + x57;
-Tbase Syntax.TZ x59 = x51 + x58;
-Tbase Syntax.TZ x60 = x50 + x59;
-Tbase Syntax.TZ x61 = x49 + x60;
-Tbase Syntax.TZ x62 = x48 + x61;
-Tbase Syntax.TZ x63 = 0x13 * x62;
-Tbase Syntax.TZ x64 = x47 + x63;
-Tbase Syntax.TZ x65 = x44 + x64;
-Tbase Syntax.TZ x66 = x65 >> 0x19;
-Tbase Syntax.TZ x67 = x6 * x18;
-Tbase Syntax.TZ x68 = x17 * 0x2;
-Tbase Syntax.TZ x69 = x7 * x68;
-Tbase Syntax.TZ x70 = x8 * x16;
-Tbase Syntax.TZ x71 = x69 + x70;
-Tbase Syntax.TZ x72 = x67 + x71;
-Tbase Syntax.TZ x73 = x15 * 0x2;
-Tbase Syntax.TZ x74 = x * x73;
-Tbase Syntax.TZ x75 = x0 * x14;
-Tbase Syntax.TZ x76 = x13 * 0x2;
-Tbase Syntax.TZ x77 = x1 * x76;
-Tbase Syntax.TZ x78 = x2 * x12;
-Tbase Syntax.TZ x79 = x11 * 0x2;
-Tbase Syntax.TZ x80 = x3 * x79;
-Tbase Syntax.TZ x81 = x4 * x10;
-Tbase Syntax.TZ x82 = x9 * 0x2;
-Tbase Syntax.TZ x83 = x5 * x82;
-Tbase Syntax.TZ x84 = x81 + x83;
-Tbase Syntax.TZ x85 = x80 + x84;
-Tbase Syntax.TZ x86 = x78 + x85;
-Tbase Syntax.TZ x87 = x77 + x86;
-Tbase Syntax.TZ x88 = x75 + x87;
-Tbase Syntax.TZ x89 = x74 + x88;
-Tbase Syntax.TZ x90 = 0x13 * x89;
-Tbase Syntax.TZ x91 = x72 + x90;
-Tbase Syntax.TZ x92 = x66 + x91;
-Tbase Syntax.TZ x93 = x92 >> 0x1a;
-Tbase Syntax.TZ x94 = x5 * x18;
-Tbase Syntax.TZ x95 = x6 * x17;
-Tbase Syntax.TZ x96 = x7 * x16;
-Tbase Syntax.TZ x97 = x8 * x15;
-Tbase Syntax.TZ x98 = x96 + x97;
-Tbase Syntax.TZ x99 = x95 + x98;
-Tbase Syntax.TZ x100 = x94 + x99;
-Tbase Syntax.TZ x101 = x * x14;
-Tbase Syntax.TZ x102 = x0 * x13;
-Tbase Syntax.TZ x103 = x1 * x12;
-Tbase Syntax.TZ x104 = x2 * x11;
-Tbase Syntax.TZ x105 = x3 * x10;
-Tbase Syntax.TZ x106 = x4 * x9;
-Tbase Syntax.TZ x107 = x105 + x106;
-Tbase Syntax.TZ x108 = x104 + x107;
-Tbase Syntax.TZ x109 = x103 + x108;
-Tbase Syntax.TZ x110 = x102 + x109;
-Tbase Syntax.TZ x111 = x101 + x110;
-Tbase Syntax.TZ x112 = 0x13 * x111;
-Tbase Syntax.TZ x113 = x100 + x112;
-Tbase Syntax.TZ x114 = x93 + x113;
-Tbase Syntax.TZ x115 = x114 >> 0x19;
-Tbase Syntax.TZ x116 = x4 * x18;
-Tbase Syntax.TZ x117 = x17 * 0x2;
-Tbase Syntax.TZ x118 = x5 * x117;
-Tbase Syntax.TZ x119 = x6 * x16;
-Tbase Syntax.TZ x120 = x15 * 0x2;
-Tbase Syntax.TZ x121 = x7 * x120;
-Tbase Syntax.TZ x122 = x8 * x14;
-Tbase Syntax.TZ x123 = x121 + x122;
-Tbase Syntax.TZ x124 = x119 + x123;
-Tbase Syntax.TZ x125 = x118 + x124;
-Tbase Syntax.TZ x126 = x116 + x125;
-Tbase Syntax.TZ x127 = x13 * 0x2;
-Tbase Syntax.TZ x128 = x * x127;
-Tbase Syntax.TZ x129 = x0 * x12;
-Tbase Syntax.TZ x130 = x11 * 0x2;
-Tbase Syntax.TZ x131 = x1 * x130;
-Tbase Syntax.TZ x132 = x2 * x10;
-Tbase Syntax.TZ x133 = x9 * 0x2;
-Tbase Syntax.TZ x134 = x3 * x133;
-Tbase Syntax.TZ x135 = x132 + x134;
-Tbase Syntax.TZ x136 = x131 + x135;
-Tbase Syntax.TZ x137 = x129 + x136;
-Tbase Syntax.TZ x138 = x128 + x137;
-Tbase Syntax.TZ x139 = 0x13 * x138;
-Tbase Syntax.TZ x140 = x126 + x139;
-Tbase Syntax.TZ x141 = x115 + x140;
-Tbase Syntax.TZ x142 = x141 >> 0x1a;
-Tbase Syntax.TZ x143 = x3 * x18;
-Tbase Syntax.TZ x144 = x4 * x17;
-Tbase Syntax.TZ x145 = x5 * x16;
-Tbase Syntax.TZ x146 = x6 * x15;
-Tbase Syntax.TZ x147 = x7 * x14;
-Tbase Syntax.TZ x148 = x8 * x13;
-Tbase Syntax.TZ x149 = x147 + x148;
-Tbase Syntax.TZ x150 = x146 + x149;
-Tbase Syntax.TZ x151 = x145 + x150;
-Tbase Syntax.TZ x152 = x144 + x151;
-Tbase Syntax.TZ x153 = x143 + x152;
-Tbase Syntax.TZ x154 = x * x12;
-Tbase Syntax.TZ x155 = x0 * x11;
-Tbase Syntax.TZ x156 = x1 * x10;
-Tbase Syntax.TZ x157 = x2 * x9;
-Tbase Syntax.TZ x158 = x156 + x157;
-Tbase Syntax.TZ x159 = x155 + x158;
-Tbase Syntax.TZ x160 = x154 + x159;
-Tbase Syntax.TZ x161 = 0x13 * x160;
-Tbase Syntax.TZ x162 = x153 + x161;
-Tbase Syntax.TZ x163 = x142 + x162;
-Tbase Syntax.TZ x164 = x163 >> 0x19;
-Tbase Syntax.TZ x165 = x2 * x18;
-Tbase Syntax.TZ x166 = x17 * 0x2;
-Tbase Syntax.TZ x167 = x3 * x166;
-Tbase Syntax.TZ x168 = x4 * x16;
-Tbase Syntax.TZ x169 = x15 * 0x2;
-Tbase Syntax.TZ x170 = x5 * x169;
-Tbase Syntax.TZ x171 = x6 * x14;
-Tbase Syntax.TZ x172 = x13 * 0x2;
-Tbase Syntax.TZ x173 = x7 * x172;
-Tbase Syntax.TZ x174 = x8 * x12;
-Tbase Syntax.TZ x175 = x173 + x174;
-Tbase Syntax.TZ x176 = x171 + x175;
-Tbase Syntax.TZ x177 = x170 + x176;
-Tbase Syntax.TZ x178 = x168 + x177;
-Tbase Syntax.TZ x179 = x167 + x178;
-Tbase Syntax.TZ x180 = x165 + x179;
-Tbase Syntax.TZ x181 = x11 * 0x2;
-Tbase Syntax.TZ x182 = x * x181;
-Tbase Syntax.TZ x183 = x0 * x10;
-Tbase Syntax.TZ x184 = x9 * 0x2;
-Tbase Syntax.TZ x185 = x1 * x184;
-Tbase Syntax.TZ x186 = x183 + x185;
-Tbase Syntax.TZ x187 = x182 + x186;
-Tbase Syntax.TZ x188 = 0x13 * x187;
-Tbase Syntax.TZ x189 = x180 + x188;
-Tbase Syntax.TZ x190 = x164 + x189;
-Tbase Syntax.TZ x191 = x190 >> 0x1a;
-Tbase Syntax.TZ x192 = x1 * x18;
-Tbase Syntax.TZ x193 = x2 * x17;
-Tbase Syntax.TZ x194 = x3 * x16;
-Tbase Syntax.TZ x195 = x4 * x15;
-Tbase Syntax.TZ x196 = x5 * x14;
-Tbase Syntax.TZ x197 = x6 * x13;
-Tbase Syntax.TZ x198 = x7 * x12;
-Tbase Syntax.TZ x199 = x8 * x11;
-Tbase Syntax.TZ x200 = x198 + x199;
-Tbase Syntax.TZ x201 = x197 + x200;
-Tbase Syntax.TZ x202 = x196 + x201;
-Tbase Syntax.TZ x203 = x195 + x202;
-Tbase Syntax.TZ x204 = x194 + x203;
-Tbase Syntax.TZ x205 = x193 + x204;
-Tbase Syntax.TZ x206 = x192 + x205;
-Tbase Syntax.TZ x207 = x * x10;
-Tbase Syntax.TZ x208 = x0 * x9;
-Tbase Syntax.TZ x209 = x207 + x208;
-Tbase Syntax.TZ x210 = 0x13 * x209;
-Tbase Syntax.TZ x211 = x206 + x210;
-Tbase Syntax.TZ x212 = x191 + x211;
-Tbase Syntax.TZ x213 = x212 >> 0x19;
-Tbase Syntax.TZ x214 = x0 * x18;
-Tbase Syntax.TZ x215 = x17 * 0x2;
-Tbase Syntax.TZ x216 = x1 * x215;
-Tbase Syntax.TZ x217 = x2 * x16;
-Tbase Syntax.TZ x218 = x15 * 0x2;
-Tbase Syntax.TZ x219 = x3 * x218;
-Tbase Syntax.TZ x220 = x4 * x14;
-Tbase Syntax.TZ x221 = x13 * 0x2;
-Tbase Syntax.TZ x222 = x5 * x221;
-Tbase Syntax.TZ x223 = x6 * x12;
-Tbase Syntax.TZ x224 = x11 * 0x2;
-Tbase Syntax.TZ x225 = x7 * x224;
-Tbase Syntax.TZ x226 = x8 * x10;
-Tbase Syntax.TZ x227 = x225 + x226;
-Tbase Syntax.TZ x228 = x223 + x227;
-Tbase Syntax.TZ x229 = x222 + x228;
-Tbase Syntax.TZ x230 = x220 + x229;
-Tbase Syntax.TZ x231 = x219 + x230;
-Tbase Syntax.TZ x232 = x217 + x231;
-Tbase Syntax.TZ x233 = x216 + x232;
-Tbase Syntax.TZ x234 = x214 + x233;
-Tbase Syntax.TZ x235 = x9 * 0x2;
-Tbase Syntax.TZ x236 = x * x235;
-Tbase Syntax.TZ x237 = 0x13 * x236;
-Tbase Syntax.TZ x238 = x234 + x237;
-Tbase Syntax.TZ x239 = x213 + x238;
-Tbase Syntax.TZ x240 = x239 >> 0x1a;
-Tbase Syntax.TZ x241 = x * x18;
-Tbase Syntax.TZ x242 = x0 * x17;
-Tbase Syntax.TZ x243 = x1 * x16;
-Tbase Syntax.TZ x244 = x2 * x15;
-Tbase Syntax.TZ x245 = x3 * x14;
-Tbase Syntax.TZ x246 = x4 * x13;
-Tbase Syntax.TZ x247 = x5 * x12;
-Tbase Syntax.TZ x248 = x6 * x11;
-Tbase Syntax.TZ x249 = x7 * x10;
-Tbase Syntax.TZ x250 = x8 * x9;
-Tbase Syntax.TZ x251 = x249 + x250;
-Tbase Syntax.TZ x252 = x248 + x251;
-Tbase Syntax.TZ x253 = x247 + x252;
-Tbase Syntax.TZ x254 = x246 + x253;
-Tbase Syntax.TZ x255 = x245 + x254;
-Tbase Syntax.TZ x256 = x244 + x255;
-Tbase Syntax.TZ x257 = x243 + x256;
-Tbase Syntax.TZ x258 = x242 + x257;
-Tbase Syntax.TZ x259 = x241 + x258;
-Tbase Syntax.TZ x260 = x240 + x259;
-Tbase Syntax.TZ x261 = x260 >> 0x19;
-Tbase Syntax.TZ x262 = 0x13 * x261;
-Tbase Syntax.TZ x263 = x43 & 0x3ffffff;
-Tbase Syntax.TZ x264 = x262 + x263;
-Tbase Syntax.TZ x265 = x264 >> 0x1a;
-Tbase Syntax.TZ x266 = x65 & 0x1ffffff;
-Tbase Syntax.TZ x267 = x265 + x266;
-Tbase Syntax.TZ x268 = x260 & 0x1ffffff;
-Tbase Syntax.TZ x269 = x239 & 0x3ffffff;
-Tbase Syntax.TZ x270 = x212 & 0x1ffffff;
-Tbase Syntax.TZ x271 = x190 & 0x3ffffff;
-Tbase Syntax.TZ x272 = x163 & 0x1ffffff;
-Tbase Syntax.TZ x273 = x141 & 0x3ffffff;
-Tbase Syntax.TZ x274 = x114 & 0x1ffffff;
-Tbase Syntax.TZ x275 = x267 >> 0x19;
-Tbase Syntax.TZ x276 = x92 & 0x3ffffff;
-Tbase Syntax.TZ x277 = x275 + x276;
-Tbase Syntax.TZ x278 = x267 & 0x1ffffff;
-Tbase Syntax.TZ x279 = x264 & 0x3ffffff;
-(Return x268, Return x269, Return x270, Return x271,
-Return x272, Return x273, Return x274, Return x277,
-Return x278, Return x279)
- : forall var : Syntax.base_type -> Type,
- expr Syntax.base_type Syntax.op
- (Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ)
-
-Argument scope is [function_scope]
diff --git a/src/Specific/GF25519Reflective/Reified/MulDisplay.v b/src/Specific/GF25519Reflective/Reified/MulDisplay.v
deleted file mode 100644
index 8ab2f65dc..000000000
--- a/src/Specific/GF25519Reflective/Reified/MulDisplay.v
+++ /dev/null
@@ -1,4 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.Reified.Mul.
-Require Export Crypto.Reflection.Z.CNotations.
-
-Print rmulW.
diff --git a/src/Specific/GF25519Reflective/Reified/MulJavaDisplay.log b/src/Specific/GF25519Reflective/Reified/MulJavaDisplay.log
deleted file mode 100644
index d0d6c01b1..000000000
--- a/src/Specific/GF25519Reflective/Reified/MulJavaDisplay.log
+++ /dev/null
@@ -1,297 +0,0 @@
-rmulW =
-fun var : Syntax.base_type -> Type =>
-x x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17
- x18 : var Syntax.TZ,
-Tbase Syntax.TZ x19 = x8 * x18;
-Tbase Syntax.TZ x20 = x17 * 0x2;
-Tbase Syntax.TZ x21 = x * x20;
-Tbase Syntax.TZ x22 = x0 * x16;
-Tbase Syntax.TZ x23 = x15 * 0x2;
-Tbase Syntax.TZ x24 = x1 * x23;
-Tbase Syntax.TZ x25 = x2 * x14;
-Tbase Syntax.TZ x26 = x13 * 0x2;
-Tbase Syntax.TZ x27 = x3 * x26;
-Tbase Syntax.TZ x28 = x4 * x12;
-Tbase Syntax.TZ x29 = x11 * 0x2;
-Tbase Syntax.TZ x30 = x5 * x29;
-Tbase Syntax.TZ x31 = x6 * x10;
-Tbase Syntax.TZ x32 = x9 * 0x2;
-Tbase Syntax.TZ x33 = x7 * x32;
-Tbase Syntax.TZ x34 = x31 + x33;
-Tbase Syntax.TZ x35 = x30 + x34;
-Tbase Syntax.TZ x36 = x28 + x35;
-Tbase Syntax.TZ x37 = x27 + x36;
-Tbase Syntax.TZ x38 = x25 + x37;
-Tbase Syntax.TZ x39 = x24 + x38;
-Tbase Syntax.TZ x40 = x22 + x39;
-Tbase Syntax.TZ x41 = x21 + x40;
-Tbase Syntax.TZ x42 = 0x13 * x41;
-Tbase Syntax.TZ x43 = x19 + x42;
-Tbase Syntax.TZ x44 = x43 >>> 0x1a;
-Tbase Syntax.TZ x45 = x7 * x18;
-Tbase Syntax.TZ x46 = x8 * x17;
-Tbase Syntax.TZ x47 = x45 + x46;
-Tbase Syntax.TZ x48 = x * x16;
-Tbase Syntax.TZ x49 = x0 * x15;
-Tbase Syntax.TZ x50 = x1 * x14;
-Tbase Syntax.TZ x51 = x2 * x13;
-Tbase Syntax.TZ x52 = x3 * x12;
-Tbase Syntax.TZ x53 = x4 * x11;
-Tbase Syntax.TZ x54 = x5 * x10;
-Tbase Syntax.TZ x55 = x6 * x9;
-Tbase Syntax.TZ x56 = x54 + x55;
-Tbase Syntax.TZ x57 = x53 + x56;
-Tbase Syntax.TZ x58 = x52 + x57;
-Tbase Syntax.TZ x59 = x51 + x58;
-Tbase Syntax.TZ x60 = x50 + x59;
-Tbase Syntax.TZ x61 = x49 + x60;
-Tbase Syntax.TZ x62 = x48 + x61;
-Tbase Syntax.TZ x63 = 0x13 * x62;
-Tbase Syntax.TZ x64 = x47 + x63;
-Tbase Syntax.TZ x65 = x44 + x64;
-Tbase Syntax.TZ x66 = x65 >>> 0x19;
-Tbase Syntax.TZ x67 = x6 * x18;
-Tbase Syntax.TZ x68 = x17 * 0x2;
-Tbase Syntax.TZ x69 = x7 * x68;
-Tbase Syntax.TZ x70 = x8 * x16;
-Tbase Syntax.TZ x71 = x69 + x70;
-Tbase Syntax.TZ x72 = x67 + x71;
-Tbase Syntax.TZ x73 = x15 * 0x2;
-Tbase Syntax.TZ x74 = x * x73;
-Tbase Syntax.TZ x75 = x0 * x14;
-Tbase Syntax.TZ x76 = x13 * 0x2;
-Tbase Syntax.TZ x77 = x1 * x76;
-Tbase Syntax.TZ x78 = x2 * x12;
-Tbase Syntax.TZ x79 = x11 * 0x2;
-Tbase Syntax.TZ x80 = x3 * x79;
-Tbase Syntax.TZ x81 = x4 * x10;
-Tbase Syntax.TZ x82 = x9 * 0x2;
-Tbase Syntax.TZ x83 = x5 * x82;
-Tbase Syntax.TZ x84 = x81 + x83;
-Tbase Syntax.TZ x85 = x80 + x84;
-Tbase Syntax.TZ x86 = x78 + x85;
-Tbase Syntax.TZ x87 = x77 + x86;
-Tbase Syntax.TZ x88 = x75 + x87;
-Tbase Syntax.TZ x89 = x74 + x88;
-Tbase Syntax.TZ x90 = 0x13 * x89;
-Tbase Syntax.TZ x91 = x72 + x90;
-Tbase Syntax.TZ x92 = x66 + x91;
-Tbase Syntax.TZ x93 = x92 >>> 0x1a;
-Tbase Syntax.TZ x94 = x5 * x18;
-Tbase Syntax.TZ x95 = x6 * x17;
-Tbase Syntax.TZ x96 = x7 * x16;
-Tbase Syntax.TZ x97 = x8 * x15;
-Tbase Syntax.TZ x98 = x96 + x97;
-Tbase Syntax.TZ x99 = x95 + x98;
-Tbase Syntax.TZ x100 = x94 + x99;
-Tbase Syntax.TZ x101 = x * x14;
-Tbase Syntax.TZ x102 = x0 * x13;
-Tbase Syntax.TZ x103 = x1 * x12;
-Tbase Syntax.TZ x104 = x2 * x11;
-Tbase Syntax.TZ x105 = x3 * x10;
-Tbase Syntax.TZ x106 = x4 * x9;
-Tbase Syntax.TZ x107 = x105 + x106;
-Tbase Syntax.TZ x108 = x104 + x107;
-Tbase Syntax.TZ x109 = x103 + x108;
-Tbase Syntax.TZ x110 = x102 + x109;
-Tbase Syntax.TZ x111 = x101 + x110;
-Tbase Syntax.TZ x112 = 0x13 * x111;
-Tbase Syntax.TZ x113 = x100 + x112;
-Tbase Syntax.TZ x114 = x93 + x113;
-Tbase Syntax.TZ x115 = x114 >>> 0x19;
-Tbase Syntax.TZ x116 = x4 * x18;
-Tbase Syntax.TZ x117 = x17 * 0x2;
-Tbase Syntax.TZ x118 = x5 * x117;
-Tbase Syntax.TZ x119 = x6 * x16;
-Tbase Syntax.TZ x120 = x15 * 0x2;
-Tbase Syntax.TZ x121 = x7 * x120;
-Tbase Syntax.TZ x122 = x8 * x14;
-Tbase Syntax.TZ x123 = x121 + x122;
-Tbase Syntax.TZ x124 = x119 + x123;
-Tbase Syntax.TZ x125 = x118 + x124;
-Tbase Syntax.TZ x126 = x116 + x125;
-Tbase Syntax.TZ x127 = x13 * 0x2;
-Tbase Syntax.TZ x128 = x * x127;
-Tbase Syntax.TZ x129 = x0 * x12;
-Tbase Syntax.TZ x130 = x11 * 0x2;
-Tbase Syntax.TZ x131 = x1 * x130;
-Tbase Syntax.TZ x132 = x2 * x10;
-Tbase Syntax.TZ x133 = x9 * 0x2;
-Tbase Syntax.TZ x134 = x3 * x133;
-Tbase Syntax.TZ x135 = x132 + x134;
-Tbase Syntax.TZ x136 = x131 + x135;
-Tbase Syntax.TZ x137 = x129 + x136;
-Tbase Syntax.TZ x138 = x128 + x137;
-Tbase Syntax.TZ x139 = 0x13 * x138;
-Tbase Syntax.TZ x140 = x126 + x139;
-Tbase Syntax.TZ x141 = x115 + x140;
-Tbase Syntax.TZ x142 = x141 >>> 0x1a;
-Tbase Syntax.TZ x143 = x3 * x18;
-Tbase Syntax.TZ x144 = x4 * x17;
-Tbase Syntax.TZ x145 = x5 * x16;
-Tbase Syntax.TZ x146 = x6 * x15;
-Tbase Syntax.TZ x147 = x7 * x14;
-Tbase Syntax.TZ x148 = x8 * x13;
-Tbase Syntax.TZ x149 = x147 + x148;
-Tbase Syntax.TZ x150 = x146 + x149;
-Tbase Syntax.TZ x151 = x145 + x150;
-Tbase Syntax.TZ x152 = x144 + x151;
-Tbase Syntax.TZ x153 = x143 + x152;
-Tbase Syntax.TZ x154 = x * x12;
-Tbase Syntax.TZ x155 = x0 * x11;
-Tbase Syntax.TZ x156 = x1 * x10;
-Tbase Syntax.TZ x157 = x2 * x9;
-Tbase Syntax.TZ x158 = x156 + x157;
-Tbase Syntax.TZ x159 = x155 + x158;
-Tbase Syntax.TZ x160 = x154 + x159;
-Tbase Syntax.TZ x161 = 0x13 * x160;
-Tbase Syntax.TZ x162 = x153 + x161;
-Tbase Syntax.TZ x163 = x142 + x162;
-Tbase Syntax.TZ x164 = x163 >>> 0x19;
-Tbase Syntax.TZ x165 = x2 * x18;
-Tbase Syntax.TZ x166 = x17 * 0x2;
-Tbase Syntax.TZ x167 = x3 * x166;
-Tbase Syntax.TZ x168 = x4 * x16;
-Tbase Syntax.TZ x169 = x15 * 0x2;
-Tbase Syntax.TZ x170 = x5 * x169;
-Tbase Syntax.TZ x171 = x6 * x14;
-Tbase Syntax.TZ x172 = x13 * 0x2;
-Tbase Syntax.TZ x173 = x7 * x172;
-Tbase Syntax.TZ x174 = x8 * x12;
-Tbase Syntax.TZ x175 = x173 + x174;
-Tbase Syntax.TZ x176 = x171 + x175;
-Tbase Syntax.TZ x177 = x170 + x176;
-Tbase Syntax.TZ x178 = x168 + x177;
-Tbase Syntax.TZ x179 = x167 + x178;
-Tbase Syntax.TZ x180 = x165 + x179;
-Tbase Syntax.TZ x181 = x11 * 0x2;
-Tbase Syntax.TZ x182 = x * x181;
-Tbase Syntax.TZ x183 = x0 * x10;
-Tbase Syntax.TZ x184 = x9 * 0x2;
-Tbase Syntax.TZ x185 = x1 * x184;
-Tbase Syntax.TZ x186 = x183 + x185;
-Tbase Syntax.TZ x187 = x182 + x186;
-Tbase Syntax.TZ x188 = 0x13 * x187;
-Tbase Syntax.TZ x189 = x180 + x188;
-Tbase Syntax.TZ x190 = x164 + x189;
-Tbase Syntax.TZ x191 = x190 >>> 0x1a;
-Tbase Syntax.TZ x192 = x1 * x18;
-Tbase Syntax.TZ x193 = x2 * x17;
-Tbase Syntax.TZ x194 = x3 * x16;
-Tbase Syntax.TZ x195 = x4 * x15;
-Tbase Syntax.TZ x196 = x5 * x14;
-Tbase Syntax.TZ x197 = x6 * x13;
-Tbase Syntax.TZ x198 = x7 * x12;
-Tbase Syntax.TZ x199 = x8 * x11;
-Tbase Syntax.TZ x200 = x198 + x199;
-Tbase Syntax.TZ x201 = x197 + x200;
-Tbase Syntax.TZ x202 = x196 + x201;
-Tbase Syntax.TZ x203 = x195 + x202;
-Tbase Syntax.TZ x204 = x194 + x203;
-Tbase Syntax.TZ x205 = x193 + x204;
-Tbase Syntax.TZ x206 = x192 + x205;
-Tbase Syntax.TZ x207 = x * x10;
-Tbase Syntax.TZ x208 = x0 * x9;
-Tbase Syntax.TZ x209 = x207 + x208;
-Tbase Syntax.TZ x210 = 0x13 * x209;
-Tbase Syntax.TZ x211 = x206 + x210;
-Tbase Syntax.TZ x212 = x191 + x211;
-Tbase Syntax.TZ x213 = x212 >>> 0x19;
-Tbase Syntax.TZ x214 = x0 * x18;
-Tbase Syntax.TZ x215 = x17 * 0x2;
-Tbase Syntax.TZ x216 = x1 * x215;
-Tbase Syntax.TZ x217 = x2 * x16;
-Tbase Syntax.TZ x218 = x15 * 0x2;
-Tbase Syntax.TZ x219 = x3 * x218;
-Tbase Syntax.TZ x220 = x4 * x14;
-Tbase Syntax.TZ x221 = x13 * 0x2;
-Tbase Syntax.TZ x222 = x5 * x221;
-Tbase Syntax.TZ x223 = x6 * x12;
-Tbase Syntax.TZ x224 = x11 * 0x2;
-Tbase Syntax.TZ x225 = x7 * x224;
-Tbase Syntax.TZ x226 = x8 * x10;
-Tbase Syntax.TZ x227 = x225 + x226;
-Tbase Syntax.TZ x228 = x223 + x227;
-Tbase Syntax.TZ x229 = x222 + x228;
-Tbase Syntax.TZ x230 = x220 + x229;
-Tbase Syntax.TZ x231 = x219 + x230;
-Tbase Syntax.TZ x232 = x217 + x231;
-Tbase Syntax.TZ x233 = x216 + x232;
-Tbase Syntax.TZ x234 = x214 + x233;
-Tbase Syntax.TZ x235 = x9 * 0x2;
-Tbase Syntax.TZ x236 = x * x235;
-Tbase Syntax.TZ x237 = 0x13 * x236;
-Tbase Syntax.TZ x238 = x234 + x237;
-Tbase Syntax.TZ x239 = x213 + x238;
-Tbase Syntax.TZ x240 = x239 >>> 0x1a;
-Tbase Syntax.TZ x241 = x * x18;
-Tbase Syntax.TZ x242 = x0 * x17;
-Tbase Syntax.TZ x243 = x1 * x16;
-Tbase Syntax.TZ x244 = x2 * x15;
-Tbase Syntax.TZ x245 = x3 * x14;
-Tbase Syntax.TZ x246 = x4 * x13;
-Tbase Syntax.TZ x247 = x5 * x12;
-Tbase Syntax.TZ x248 = x6 * x11;
-Tbase Syntax.TZ x249 = x7 * x10;
-Tbase Syntax.TZ x250 = x8 * x9;
-Tbase Syntax.TZ x251 = x249 + x250;
-Tbase Syntax.TZ x252 = x248 + x251;
-Tbase Syntax.TZ x253 = x247 + x252;
-Tbase Syntax.TZ x254 = x246 + x253;
-Tbase Syntax.TZ x255 = x245 + x254;
-Tbase Syntax.TZ x256 = x244 + x255;
-Tbase Syntax.TZ x257 = x243 + x256;
-Tbase Syntax.TZ x258 = x242 + x257;
-Tbase Syntax.TZ x259 = x241 + x258;
-Tbase Syntax.TZ x260 = x240 + x259;
-Tbase Syntax.TZ x261 = x260 >>> 0x19;
-Tbase Syntax.TZ x262 = 0x13 * x261;
-Tbase Syntax.TZ x263 = x43 & 0x3ffffff;
-Tbase Syntax.TZ x264 = x262 + x263;
-Tbase Syntax.TZ x265 = x264 >>> 0x1a;
-Tbase Syntax.TZ x266 = x65 & 0x1ffffff;
-Tbase Syntax.TZ x267 = x265 + x266;
-Tbase Syntax.TZ x268 = x260 & 0x1ffffff;
-Tbase Syntax.TZ x269 = x239 & 0x3ffffff;
-Tbase Syntax.TZ x270 = x212 & 0x1ffffff;
-Tbase Syntax.TZ x271 = x190 & 0x3ffffff;
-Tbase Syntax.TZ x272 = x163 & 0x1ffffff;
-Tbase Syntax.TZ x273 = x141 & 0x3ffffff;
-Tbase Syntax.TZ x274 = x114 & 0x1ffffff;
-Tbase Syntax.TZ x275 = x267 >>> 0x19;
-Tbase Syntax.TZ x276 = x92 & 0x3ffffff;
-Tbase Syntax.TZ x277 = x275 + x276;
-Tbase Syntax.TZ x278 = x267 & 0x1ffffff;
-Tbase Syntax.TZ x279 = x264 & 0x3ffffff;
-(Return x268, Return x269, Return x270, Return x271,
-Return x272, Return x273, Return x274, Return x277,
-Return x278, Return x279)
- : forall var : Syntax.base_type -> Type,
- expr Syntax.base_type Syntax.op
- (Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Syntax.TZ ->
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ * Tbase Syntax.TZ * Tbase Syntax.TZ *
- Tbase Syntax.TZ)
-
-Argument scope is [function_scope]
diff --git a/src/Specific/GF25519Reflective/Reified/MulJavaDisplay.v b/src/Specific/GF25519Reflective/Reified/MulJavaDisplay.v
deleted file mode 100644
index ed82f9e6c..000000000
--- a/src/Specific/GF25519Reflective/Reified/MulJavaDisplay.v
+++ /dev/null
@@ -1,4 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.Reified.Mul.
-Require Export Crypto.Reflection.Z.JavaNotations.
-
-Print rmulW.
diff --git a/src/Specific/GF25519Reflective/Reified/Opp.v b/src/Specific/GF25519Reflective/Reified/Opp.v
deleted file mode 100644
index 36a0b907e..000000000
--- a/src/Specific/GF25519Reflective/Reified/Opp.v
+++ /dev/null
@@ -1,12 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonUnOp.
-
-Definition roppZ_sig : rexpr_unop_sig opp. Proof. reify_sig. Defined.
-Definition roppW := Eval vm_compute in rword_of_Z roppZ_sig.
-Lemma roppW_correct_and_bounded_gen : correct_and_bounded_genT roppW roppZ_sig.
-Proof. rexpr_correct. Qed.
-Definition ropp_output_bounds := Eval vm_compute in compute_bounds roppW ExprUnOp_bounds.
-
-Local Open Scope string_scope.
-Compute ("Opp", compute_bounds_for_display roppW ExprUnOp_bounds).
-Compute ("Opp overflows? ", sanity_compute roppW ExprUnOp_bounds).
-Compute ("Opp overflows (error if it does)? ", sanity_check roppW ExprUnOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/Pack.v b/src/Specific/GF25519Reflective/Reified/Pack.v
deleted file mode 100644
index 9645cf24c..000000000
--- a/src/Specific/GF25519Reflective/Reified/Pack.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonUnOpFEToWire.
-
-Definition rpackZ_sig : rexpr_unop_FEToWire_sig pack. Proof. reify_sig. Defined.
-Definition rpackW := Eval vm_compute in rword_of_Z rpackZ_sig.
-Lemma rpackW_correct_and_bounded_gen : correct_and_bounded_genT rpackW rpackZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rpack_output_bounds := Eval vm_compute in compute_bounds rpackW ExprUnOpFEToWire_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition rpackW_correct_and_bounded
- := ExprUnOpFEToWire_correct_and_bounded
- rpackW pack rpackZ_sig rpackW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("Pack", compute_bounds_for_display rpackW ExprUnOpFEToWire_bounds).
-Compute ("Pack overflows? ", sanity_compute rpackW ExprUnOpFEToWire_bounds).
-Compute ("Pack overflows (error if it does)? ", sanity_check rpackW ExprUnOpFEToWire_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/PreFreeze.v b/src/Specific/GF25519Reflective/Reified/PreFreeze.v
deleted file mode 100644
index 6b3bce924..000000000
--- a/src/Specific/GF25519Reflective/Reified/PreFreeze.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonUnOp.
-
-Definition rprefreezeZ_sig : rexpr_unop_sig prefreeze. Proof. reify_sig. Defined.
-Definition rprefreezeW := Eval vm_compute in rword_of_Z rprefreezeZ_sig.
-Lemma rprefreezeW_correct_and_bounded_gen : correct_and_bounded_genT rprefreezeW rprefreezeZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rprefreeze_output_bounds := Eval vm_compute in compute_bounds rprefreezeW ExprUnOp_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition rprefreezeW_correct_and_bounded
- := ExprUnOp_correct_and_bounded
- rprefreezeW prefreeze rprefreezeZ_sig rprefreezeW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("PreFreeze", compute_bounds_for_display rprefreezeW ExprUnOp_bounds).
-Compute ("PreFreeze overflows? ", sanity_compute rprefreezeW ExprUnOp_bounds).
-Compute ("PreFreeze overflows (error if it does)? ", sanity_check rprefreezeW ExprUnOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/Sub.v b/src/Specific/GF25519Reflective/Reified/Sub.v
deleted file mode 100644
index aae3af64b..000000000
--- a/src/Specific/GF25519Reflective/Reified/Sub.v
+++ /dev/null
@@ -1,12 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonBinOp.
-
-Definition rsubZ_sig : rexpr_binop_sig sub. Proof. reify_sig. Defined.
-Definition rsubW := Eval vm_compute in rword_of_Z rsubZ_sig.
-Lemma rsubW_correct_and_bounded_gen : correct_and_bounded_genT rsubW rsubZ_sig.
-Proof. rexpr_correct. Qed.
-Definition rsub_output_bounds := Eval vm_compute in compute_bounds rsubW ExprBinOp_bounds.
-
-Local Open Scope string_scope.
-Compute ("Sub", compute_bounds_for_display rsubW ExprBinOp_bounds).
-Compute ("Sub overflows? ", sanity_compute rsubW ExprBinOp_bounds).
-Compute ("Sub overflows (error if it does)? ", sanity_check rsubW ExprBinOp_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/Unpack.v b/src/Specific/GF25519Reflective/Reified/Unpack.v
deleted file mode 100644
index f8ab40383..000000000
--- a/src/Specific/GF25519Reflective/Reified/Unpack.v
+++ /dev/null
@@ -1,17 +0,0 @@
-Require Import Crypto.Specific.GF25519Reflective.CommonUnOpWireToFE.
-
-Definition runpackZ_sig : rexpr_unop_WireToFE_sig unpack. Proof. reify_sig. Defined.
-Definition runpackW := Eval vm_compute in rword_of_Z runpackZ_sig.
-Lemma runpackW_correct_and_bounded_gen : correct_and_bounded_genT runpackW runpackZ_sig.
-Proof. rexpr_correct. Qed.
-Definition runpack_output_bounds := Eval vm_compute in compute_bounds runpackW ExprUnOpWireToFE_bounds.
-Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition runpackW_correct_and_bounded
- := ExprUnOpWireToFE_correct_and_bounded
- runpackW unpack runpackZ_sig runpackW_correct_and_bounded_gen
- _ _.
-
-Local Open Scope string_scope.
-Compute ("Unpack", compute_bounds_for_display runpackW ExprUnOpWireToFE_bounds).
-Compute ("Unpack overflows? ", sanity_compute runpackW ExprUnOpWireToFE_bounds).
-Compute ("Unpack overflows (error if it does)? ", sanity_check runpackW ExprUnOpWireToFE_bounds).
diff --git a/src/Specific/GF25519Reflective/Reified/rebuild-reified.py b/src/Specific/GF25519Reflective/Reified/rebuild-reified.py
deleted file mode 100755
index f9348d0cc..000000000
--- a/src/Specific/GF25519Reflective/Reified/rebuild-reified.py
+++ /dev/null
@@ -1,30 +0,0 @@
-#!/usr/bin/env python2.7
-from __future__ import with_statement
-
-for name, opkind in ([(name, 'BinOp') for name in ('Add', 'Carry_Add', 'Sub', 'Carry_Sub', 'Mul')]
- + [(name, 'UnOp') for name in ('Opp', 'Carry_Opp', 'PreFreeze')]
- + [('Ge_Modulus', 'UnOp_FEToZ'), ('Pack', 'UnOp_FEToWire'), ('Unpack', 'UnOp_WireToFE')]):
- lname = name.lower()
- lopkind = opkind.replace('UnOp', 'unop').replace('BinOp', 'binop')
- uopkind = opkind.replace('_', '')
- extra = ''
- #if name in ('Carry_Add', 'Carry_Sub', 'Mul', 'Carry_Opp', 'Pack', 'Unpack', 'Ge_Modulus', 'PreFreeze'):
- extra = r"""Local Obligation Tactic := intros; vm_compute; constructor.
-Program Definition r%(lname)sZ_correct_and_bounded
- := rexpr_correct_and_bounded r%(lname)sZ r%(lname)sZ_Wf Expr%(uopkind)s_bounds.
-""" % locals()
- with open(name.replace('_', '') + '.v', 'w') as f:
- f.write(r"""Require Import Crypto.Specific.GF25519Reflective.Common.
-
-Definition r%(lname)sZ_sig : rexpr_%(lopkind)s_sig %(lname)s. Proof. reify_sig. Defined.
-Definition r%(lname)sZ : Syntax.Expr _ _ _ := Eval vm_compute in proj1_sig r%(lname)sZ_sig.
-Definition r%(lname)sW : Syntax.Expr _ _ _ := Eval vm_compute in rexpr_select_word_sizes r%(lname)sZ Expr%(uopkind)s_bounds.
-Definition r%(lname)sZ_Wf : rexpr_wfT r%(lname)sZ. Proof. prove_rexpr_wfT. Qed.
-Definition r%(lname)s_output_bounds
- := Eval vm_compute in compute_bounds r%(lname)sZ Expr%(uopkind)s_bounds.
-%(extra)s
-Local Open Scope string_scope.
-Compute ("%(name)s", compute_bounds_for_display r%(lname)sZ Expr%(uopkind)s_bounds).
-Compute ("%(name)s overflows? ", sanity_compute r%(lname)sZ Expr%(uopkind)s_bounds).
-Compute ("%(name)s overflows (error if it does)? ", sanity_check r%(lname)sZ Expr%(uopkind)s_bounds).
-""" % locals())