aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs
diff options
context:
space:
mode:
Diffstat (limited to 'test-suite/bugs')
-rw-r--r--test-suite/bugs/closed/1784.v5
-rw-r--r--test-suite/bugs/closed/1850.v4
-rw-r--r--test-suite/bugs/closed/2016.v6
-rw-r--r--test-suite/bugs/closed/2021.v2
-rw-r--r--test-suite/bugs/closed/2800.v (renamed from test-suite/bugs/opened/2800.v)2
-rw-r--r--test-suite/bugs/closed/3080.v18
-rw-r--r--test-suite/bugs/closed/3251.v1
-rw-r--r--test-suite/bugs/closed/3612.v5
-rw-r--r--test-suite/bugs/closed/3649.v5
-rw-r--r--test-suite/bugs/closed/3690.v4
-rw-r--r--test-suite/bugs/closed/3699.v4
-rw-r--r--test-suite/bugs/closed/3849.v (renamed from test-suite/bugs/opened/3849.v)2
-rw-r--r--test-suite/bugs/closed/3881.v2
-rw-r--r--test-suite/bugs/closed/3911.v26
-rw-r--r--test-suite/bugs/closed/3929.v67
-rw-r--r--test-suite/bugs/closed/3957.v6
-rw-r--r--test-suite/bugs/closed/4214.v (renamed from test-suite/bugs/opened/4214.v)3
-rw-r--r--test-suite/bugs/closed/4479.v3
-rw-r--r--test-suite/bugs/closed/4622.v24
-rw-r--r--test-suite/bugs/closed/4726.v19
-rw-r--r--test-suite/bugs/closed/4787.v9
-rw-r--r--test-suite/bugs/closed/4811.v1685
-rw-r--r--test-suite/bugs/closed/4816.v22
-rw-r--r--test-suite/bugs/closed/4865.v52
-rw-r--r--test-suite/bugs/closed/4873.v72
-rw-r--r--test-suite/bugs/closed/HoTT_coq_047.v2
-rw-r--r--test-suite/bugs/opened/3410.v1
-rw-r--r--test-suite/bugs/opened/3889.v11
-rw-r--r--test-suite/bugs/opened/3890.v18
-rw-r--r--test-suite/bugs/opened/3916.v3
-rw-r--r--test-suite/bugs/opened/3919.v-disabled13
-rw-r--r--test-suite/bugs/opened/3922.v-disabled83
-rw-r--r--test-suite/bugs/opened/3926.v30
-rw-r--r--test-suite/bugs/opened/3928.v-disabled12
-rw-r--r--test-suite/bugs/opened/3938.v6
-rw-r--r--test-suite/bugs/opened/3946.v11
-rw-r--r--test-suite/bugs/opened/3948.v25
-rw-r--r--test-suite/bugs/opened/4813.v5
38 files changed, 2251 insertions, 17 deletions
diff --git a/test-suite/bugs/closed/1784.v b/test-suite/bugs/closed/1784.v
index 0b63d7b56..25d1b192e 100644
--- a/test-suite/bugs/closed/1784.v
+++ b/test-suite/bugs/closed/1784.v
@@ -91,9 +91,8 @@ Next Obligation. intro H; inversion H. Defined.
Next Obligation. intro H; inversion H; subst. Defined.
Next Obligation.
intro H1; contradict H. inversion H1; subst. assumption.
- contradict H0; assumption. Defined.
-Next Obligation.
- intro H1; contradict H0. inversion H1; subst. assumption. Defined.
+ contradict H0; assumption. Defined.
+Next Obligation. intro H1; contradict H0. inversion H1; subst. assumption. Defined.
Next Obligation.
intro H1; contradict H. inversion H1; subst. assumption. Defined.
Next Obligation.
diff --git a/test-suite/bugs/closed/1850.v b/test-suite/bugs/closed/1850.v
new file mode 100644
index 000000000..26b48093b
--- /dev/null
+++ b/test-suite/bugs/closed/1850.v
@@ -0,0 +1,4 @@
+Parameter P : Type -> Type -> Type.
+Notation "e |= t --> v" := (P e t v) (at level 100, t at level 54).
+Fail Check (nat |= nat --> nat).
+
diff --git a/test-suite/bugs/closed/2016.v b/test-suite/bugs/closed/2016.v
index 13ec5bea9..536e6fabd 100644
--- a/test-suite/bugs/closed/2016.v
+++ b/test-suite/bugs/closed/2016.v
@@ -1,6 +1,8 @@
(* Coq 8.2beta4 *)
Require Import Classical_Prop.
+Unset Structural Injection.
+
Record coreSemantics : Type := CoreSemantics {
core: Type;
corestep: core -> core -> Prop;
@@ -49,7 +51,7 @@ unfold oe_corestep; intros.
assert (HH:= step_fun _ _ _ H H0); clear H H0.
destruct q1; destruct q2; unfold oe2coreSem; simpl in *.
generalize (inj_pairT1 _ _ _ _ _ _ HH); clear HH; intros.
-injection H; clear H; intros.
+injection H.
revert in_q1 in_corestep1 in_corestep_fun1
H.
pattern in_core1.
@@ -59,4 +61,4 @@ apply sym_eq.
(** good to here **)
Show Universes.
Print Universes.
-Fail apply H0. \ No newline at end of file
+Fail apply H0.
diff --git a/test-suite/bugs/closed/2021.v b/test-suite/bugs/closed/2021.v
index e598e5aed..5df92998e 100644
--- a/test-suite/bugs/closed/2021.v
+++ b/test-suite/bugs/closed/2021.v
@@ -1,6 +1,8 @@
(* correct failure of injection/discriminate on types whose inductive
status derives from the substitution of an argument *)
+Unset Structural Injection.
+
Inductive t : nat -> Type :=
| M : forall n: nat, nat -> t n.
diff --git a/test-suite/bugs/opened/2800.v b/test-suite/bugs/closed/2800.v
index c559ab0c1..2ee438934 100644
--- a/test-suite/bugs/opened/2800.v
+++ b/test-suite/bugs/closed/2800.v
@@ -1,6 +1,6 @@
Goal False.
-Fail intuition
+intuition
match goal with
| |- _ => idtac " foo"
end.
diff --git a/test-suite/bugs/closed/3080.v b/test-suite/bugs/closed/3080.v
new file mode 100644
index 000000000..7d0dc090e
--- /dev/null
+++ b/test-suite/bugs/closed/3080.v
@@ -0,0 +1,18 @@
+(* -*- coq-prog-args: ("-emacs" "-nois") -*- *)
+Delimit Scope type_scope with type.
+Delimit Scope function_scope with function.
+
+Bind Scope type_scope with Sortclass.
+Bind Scope function_scope with Funclass.
+
+Reserved Notation "x -> y" (at level 99, right associativity, y at level 200).
+Notation "A -> B" := (forall (_ : A), B) : type_scope.
+
+Definition compose {A B C} (g : B -> C) (f : A -> B) :=
+ fun x : A => g (f x).
+
+Notation " g ∘ f " := (compose g f)
+ (at level 40, left associativity) : function_scope.
+
+Fail Check (fun x => x) ∘ (fun x => x). (* this [Check] should fail, as [function_scope] is not opened *)
+Check compose ((fun x => x) ∘ (fun x => x)) (fun x => x). (* this check should succeed, as [function_scope] should be automatically bound in the arugments to [compose] *)
diff --git a/test-suite/bugs/closed/3251.v b/test-suite/bugs/closed/3251.v
index 5a7ae2002..d4ce050c5 100644
--- a/test-suite/bugs/closed/3251.v
+++ b/test-suite/bugs/closed/3251.v
@@ -1,4 +1,5 @@
Goal True.
+idtac.
Ltac foo := idtac.
(* print out happens twice:
foo is defined
diff --git a/test-suite/bugs/closed/3612.v b/test-suite/bugs/closed/3612.v
index 9125ab16d..a54768507 100644
--- a/test-suite/bugs/closed/3612.v
+++ b/test-suite/bugs/closed/3612.v
@@ -6,6 +6,8 @@ lines, then from 421 lines to 428 lines, then from 444 lines to 429 lines, then
Reserved Notation "x -> y" (at level 99, right associativity, y at level 200).
Reserved Notation "x = y :> T" (at level 70, y at next level, no associativity).
Reserved Notation "x = y" (at level 70, no associativity).
+Delimit Scope type_scope with type.
+Bind Scope type_scope with Sortclass.
Open Scope type_scope.
Global Set Universe Polymorphism.
Notation "A -> B" := (forall (_ : A), B) : type_scope.
@@ -35,6 +37,9 @@ Axiom path_path_sigma : forall {A : Type} (P : A -> Type) (u v : sigT P)
(r : p..1 = q..1)
(s : transport (fun x => transport P x u.2 = v.2) r p..2 = q..2),
p = q.
+
+Declare ML Module "coretactics".
+
Goal forall (A : Type) (B : forall _ : A, Type) (x : @sigT A (fun x : A => B x))
(xx : @paths (@sigT A (fun x0 : A => B x0)) x x),
@paths (@paths (@sigT A (fun x0 : A => B x0)) x x) xx
diff --git a/test-suite/bugs/closed/3649.v b/test-suite/bugs/closed/3649.v
index 06188e7b1..fc4c171e2 100644
--- a/test-suite/bugs/closed/3649.v
+++ b/test-suite/bugs/closed/3649.v
@@ -2,8 +2,11 @@
(* File reduced by coq-bug-finder from original input, then from 9518 lines to 404 lines, then from 410 lines to 208 lines, then from 162 lines to 77 lines *)
(* coqc version trunk (September 2014) compiled on Sep 18 2014 21:0:5 with OCaml 4.01.0
coqtop version cagnode16:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (07e4438bd758c2ced8caf09a6961ccd77d84e42b) *)
+Declare ML Module "coretactics".
Reserved Notation "x -> y" (at level 99, right associativity, y at level 200).
Reserved Notation "x = y" (at level 70, no associativity).
+Delimit Scope type_scope with type.
+Bind Scope type_scope with Sortclass.
Open Scope type_scope.
Axiom admit : forall {T}, T.
Notation "A -> B" := (forall (_ : A), B) : type_scope.
@@ -54,4 +57,4 @@ Goal forall (C D : PreCategory) (G G' : Functor C D)
(** This [change] succeeded, but did not progress, in 07e4438bd758c2ced8caf09a6961ccd77d84e42b, because [T0 x o T1 x] was not found in the goal *)
let T0 := match goal with |- context[components_of ?T0 ?x o components_of ?T1 ?x] => constr:(T0) end in
let T1 := match goal with |- context[components_of ?T0 ?x o components_of ?T1 ?x] => constr:(T1) end in
- progress change (T0 x o T1 x) with ((fun y => y) (T0 x o T1 x)). \ No newline at end of file
+ progress change (T0 x o T1 x) with ((fun y => y) (T0 x o T1 x)).
diff --git a/test-suite/bugs/closed/3690.v b/test-suite/bugs/closed/3690.v
index c24173abf..fd9640b89 100644
--- a/test-suite/bugs/closed/3690.v
+++ b/test-suite/bugs/closed/3690.v
@@ -47,7 +47,7 @@ Type@{Top.21} -> Type@{Top.23}
Top.23 < Top.22
*) *)
Fail Check @qux@{Set Set}.
-Fail Check @qux@{Set Set Set}.
+Check @qux@{Type Type Type Type}.
(* [qux] should only need two universes *)
-Check @qux@{i j k}. (* Error: The command has not failed!, but I think this is suboptimal *)
+Check @qux@{i j k l}. (* Error: The command has not failed!, but I think this is suboptimal *)
Fail Check @qux@{i j}.
diff --git a/test-suite/bugs/closed/3699.v b/test-suite/bugs/closed/3699.v
index aad0bb44d..8dadc2419 100644
--- a/test-suite/bugs/closed/3699.v
+++ b/test-suite/bugs/closed/3699.v
@@ -65,7 +65,7 @@ Module NonPrim.
set (fibermap := fun a0p : hfiber f (f a)
=> let (a0, p) := a0p in transport P p (d a0)).
Set Printing Implicit.
- let G := match goal with |- ?G => constr:G end in
+ let G := match goal with |- ?G => constr:(G) end in
first [ match goal with
| [ |- (@isconnected_elim n (@hfiber A B f (f a))
(@isconnected_hfiber_conn_map n A B f H (f a))
@@ -142,7 +142,7 @@ Module Prim.
set (fibermap := fun a0p : hfiber f (f a)
=> let (a0, p) := a0p in transport P p (d a0)).
Set Printing Implicit.
- let G := match goal with |- ?G => constr:G end in
+ let G := match goal with |- ?G => constr:(G) end in
first [ match goal with
| [ |- (@isconnected_elim n (@hfiber A B f (f a))
(@isconnected_hfiber_conn_map n A B f H (f a))
diff --git a/test-suite/bugs/opened/3849.v b/test-suite/bugs/closed/3849.v
index 5290054a0..a8dc3af9c 100644
--- a/test-suite/bugs/opened/3849.v
+++ b/test-suite/bugs/closed/3849.v
@@ -5,4 +5,4 @@ Tactic Notation "bar" hyp_list(hs) := foo hs.
Goal True.
do 5 pose proof 0 as ?n0.
foo n1 n2.
-Fail bar n3 n4.
+bar n3 n4.
diff --git a/test-suite/bugs/closed/3881.v b/test-suite/bugs/closed/3881.v
index 070d1e9c7..a327bbf2a 100644
--- a/test-suite/bugs/closed/3881.v
+++ b/test-suite/bugs/closed/3881.v
@@ -23,7 +23,7 @@ Proof.
pose (fun H => @isequiv_homotopic _ _ ((g o f) o f^-1) _ H
(fun b => ap g (eisretr f b))) as k.
revert k.
- let x := match goal with |- let k := ?x in _ => constr:x end in
+ let x := match goal with |- let k := ?x in _ => constr:(x) end in
intro k; clear k;
pose (x _).
pose (@isequiv_homotopic _ _ ((g o f) o f^-1) g _
diff --git a/test-suite/bugs/closed/3911.v b/test-suite/bugs/closed/3911.v
new file mode 100644
index 000000000..b289eafbf
--- /dev/null
+++ b/test-suite/bugs/closed/3911.v
@@ -0,0 +1,26 @@
+(* Tested against coq ee596bc *)
+
+Set Nonrecursive Elimination Schemes.
+Set Primitive Projections.
+Set Universe Polymorphism.
+
+Record setoid := { base : Type }.
+
+Definition catdata (Obj Arr : Type) : Type := nat.
+ (* [nat] can be replaced by any other type, it seems,
+ without changing the error *)
+
+Record cat : Type :=
+ {
+ obj : setoid;
+ arr : Type;
+ dta : catdata (base obj) arr
+ }.
+
+Definition bcwa (C:cat) (B:setoid) :Type := nat.
+ (* As above, nothing special about [nat] here. *)
+
+Record temp {C}{B} (e:bcwa C B) :=
+ { fld : base (obj C) }.
+
+Print temp_rect.
diff --git a/test-suite/bugs/closed/3929.v b/test-suite/bugs/closed/3929.v
new file mode 100644
index 000000000..955581ef2
--- /dev/null
+++ b/test-suite/bugs/closed/3929.v
@@ -0,0 +1,67 @@
+Universes i j.
+Set Printing Universes.
+Set Printing All.
+Polymorphic Definition lt@{x y} : Type@{y} := Type@{x}.
+Goal True.
+evar (T:Type@{i}).
+set (Z := nat : Type@{j}). simpl in Z.
+let Tv:=eval cbv [T] in T in
+pose (x:=Tv).
+revert x.
+refine (_ : let x:=Z in True).
+(** This enforces i <= j *)
+Fail pose (lt@{i j}).
+let Zv:=eval cbv [Z] in Z in
+let Tv:=eval cbv [T] in T in
+constr_eq Zv Tv.
+exact I.
+Defined.
+
+Goal True.
+evar (T:nat).
+pose (Z:=0).
+let Tv:=eval cbv [T] in T in
+pose (x:=Tv).
+revert x.
+refine (_ : let x:=Z in True).
+let Zv:=eval cbv [Z] in Z in
+let Tv:=eval cbv [T] in T in
+constr_eq Zv Tv.
+Abort.
+
+Goal True.
+evar (T:Set).
+pose (Z:=nat).
+let Tv:=eval cbv [T] in T in
+pose (x:=Tv).
+revert x.
+refine (_ : let x:=Z in True).
+let Zv:=eval cbv [Z] in Z in
+let Tv:=eval cbv [T] in T in
+constr_eq Zv Tv.
+Abort.
+
+Goal forall (A:Type)(a:A), True.
+intros A a.
+evar (T:A).
+pose (Z:=a).
+let Tv:=eval cbv delta [T] in T in
+pose (x:=Tv).
+revert x.
+refine (_ : let x:=Z in True).
+let Zv:=eval cbv [Z] in Z in
+let Tv:=eval cbv [T] in T in
+constr_eq Zv Tv.
+Abort.
+
+Goal True.
+evar (T:Type).
+pose (Z:=nat).
+let Tv:=eval cbv [T] in T in
+pose (x:=Tv).
+revert x.
+refine (_ : let x:=Z in True).
+let Zv:=eval cbv [Z] in Z in
+let Tv:=eval cbv [T] in T in
+constr_eq Zv Tv.
+Abort.
diff --git a/test-suite/bugs/closed/3957.v b/test-suite/bugs/closed/3957.v
new file mode 100644
index 000000000..e20a6e97f
--- /dev/null
+++ b/test-suite/bugs/closed/3957.v
@@ -0,0 +1,6 @@
+Ltac foo tac := tac.
+
+Goal True.
+Proof.
+foo subst.
+Admitted.
diff --git a/test-suite/bugs/opened/4214.v b/test-suite/bugs/closed/4214.v
index 3daf45213..d684e8cf4 100644
--- a/test-suite/bugs/opened/4214.v
+++ b/test-suite/bugs/closed/4214.v
@@ -2,4 +2,5 @@
Goal forall A (a b c : A), b = a -> b = c -> a = c.
intros.
subst.
-Fail reflexivity.
+reflexivity.
+Qed. \ No newline at end of file
diff --git a/test-suite/bugs/closed/4479.v b/test-suite/bugs/closed/4479.v
new file mode 100644
index 000000000..921579d1e
--- /dev/null
+++ b/test-suite/bugs/closed/4479.v
@@ -0,0 +1,3 @@
+Goal True.
+Fail autorewrite with foo.
+try autorewrite with foo.
diff --git a/test-suite/bugs/closed/4622.v b/test-suite/bugs/closed/4622.v
new file mode 100644
index 000000000..ffa478cb8
--- /dev/null
+++ b/test-suite/bugs/closed/4622.v
@@ -0,0 +1,24 @@
+Set Primitive Projections.
+
+Record foo : Type := bar { x : unit }.
+
+Goal forall t u, bar t = bar u -> t = u.
+Proof.
+ intros.
+ injection H.
+ trivial.
+Qed.
+(* Was: Error: Pattern-matching expression on an object of inductive type foo has invalid information. *)
+
+(** Dependent pattern-matching is ok on this one as it has eta *)
+Definition baz (x : foo) :=
+ match x as x' return x' = x' with
+ | bar u => eq_refl
+ end.
+
+Inductive foo' : Type := bar' {x' : unit; y: foo'}.
+(** Dependent pattern-matching is not ok on this one *)
+Fail Definition baz' (x : foo') :=
+ match x as x' return x' = x' with
+ | bar' u y => eq_refl
+ end.
diff --git a/test-suite/bugs/closed/4726.v b/test-suite/bugs/closed/4726.v
new file mode 100644
index 000000000..0037b6fde
--- /dev/null
+++ b/test-suite/bugs/closed/4726.v
@@ -0,0 +1,19 @@
+Set Universe Polymorphism.
+
+Definition le@{i j} : Type@{j} :=
+ (fun A : Type@{j} => A)
+ (unit : Type@{i}).
+Definition eq@{i j} : Type@{j} := let x := le@{i j} in le@{j i}.
+
+Record Inj@{i j} (A : Type@{i}) (B : Type@{j}) : Type@{j} :=
+ { inj : A }.
+
+Monomorphic Universe u1.
+Let ty1 : Type@{u1} := Set.
+Check Inj@{Set u1}.
+(* Would fail with univ inconsistency if the universe was minimized *)
+
+Record Inj'@{i j} (A : Type@{i}) (B : Type@{j}) : Type@{j} :=
+ { inj' : A; foo : Type@{j} := eq@{i j} }.
+Fail Check Inj'@{Set u1}. (* Do not drop constraint i = j *)
+Check Inj'@{Set Set}.
diff --git a/test-suite/bugs/closed/4787.v b/test-suite/bugs/closed/4787.v
new file mode 100644
index 000000000..b586cba50
--- /dev/null
+++ b/test-suite/bugs/closed/4787.v
@@ -0,0 +1,9 @@
+(* [Unset Bracketing Last Introduction Pattern] was not working *)
+
+Unset Bracketing Last Introduction Pattern.
+
+Goal forall T (x y : T * T), fst x = fst y /\ snd x = snd y -> x = y.
+do 10 ((intros [] || intro); simpl); reflexivity.
+Qed.
+
+
diff --git a/test-suite/bugs/closed/4811.v b/test-suite/bugs/closed/4811.v
new file mode 100644
index 000000000..a91496262
--- /dev/null
+++ b/test-suite/bugs/closed/4811.v
@@ -0,0 +1,1685 @@
+(* Test about a slowness of f_equal in 8.5pl1 *)
+
+(* Submitted by Jason Gross *)
+
+(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "src" "Crypto" "-R" "Bedrock" "Bedrock" "-R" "coqprime-8.5/Coqprime" "Coqprime" "-top" "GF255192") -*- *)
+(* File reduced by coq-bug-finder from original input, then from 162 lines to 23 lines, then from 245 lines to 95 lines, then from 198 lines to 101 lines, then from 654 lines to 452 lines, then from 591 lines to 505 lines, then from 1770 lines to 580 lines, then from 2238 lines to 1715 lines, then from 1776 lines to 1738 lines, then from 1750 lines to 1679 lines, then from 1693 lines to 1679 lines *)
+(* coqc version 8.5pl1 (April 2016) compiled on Apr 18 2016 14:48:5 with OCaml 4.02.3
+ coqtop version 8.5pl1 (April 2016) *)
+Require Coq.ZArith.ZArith.
+
+Import Coq.ZArith.ZArith.
+
+Axiom F : Z -> Set.
+Definition Let_In {A P} (x : A) (f : forall y : A, P y)
+ := let y := x in f y.
+Local Open Scope Z_scope.
+Definition modulus : Z := 2^255 - 19.
+Axiom decode : list Z -> F modulus.
+Goal forall x9 x8 x7 x6 x5 x4 x3 x2 x1 x0 y9 y8 y7 y6 y5 y4 y3 y2 y1 y0 : Z,
+ let Zmul := Z.mul in
+ let Zadd := Z.add in
+ let Zsub := Z.sub in
+ let Zpow_pos := Z.pow_pos in
+ @eq (F (Zsub (Zpow_pos (Zpos (xO xH)) (xI (xI (xI (xI (xI (xI (xI xH)))))))) (Zpos (xI (xI (xO (xO xH)))))))
+ (@decode
+ (@Let_In Z (fun _ : Z => list Z)
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH))) (Zmul x8 y2)) (Zmul (Zmul x7 y3) (Zpos (xO xH))))
+ (Zmul x6 y4)) (Zmul (Zmul x5 y5) (Zpos (xO xH)))) (Zmul x4 y6))
+ (Zmul (Zmul x3 y7) (Zpos (xO xH)))) (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (fun z : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zadd (Zmul x9 y2) (Zmul x8 y3)) (Zmul x7 y4)) (Zmul x6 y5)) (Zmul x5 y6))
+ (Zmul x4 y7)) (Zmul x3 y8)) (Zmul x2 y9)))))
+ (fun z0 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z0 (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH)))) (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH))) (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH)))) (Zmul x6 y6)) (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8)) (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (fun z1 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z1 (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2)) (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5)) (Zmul x7 y6)) (Zmul x6 y7)) (Zmul x5 y8))
+ (Zmul x4 y9)))))
+ (fun z2 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z2 (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH)))) (Zmul x2 y2))
+ (Zmul (Zmul x1 y3) (Zpos (xO xH)))) (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH))) (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH)))) (Zmul x6 y8)) (Zmul (Zmul x5 y9) (Zpos (xO xH)))))))
+ (fun z3 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z3 (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2)) (Zmul x2 y3)) (Zmul x1 y4))
+ (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8)) (Zmul x6 y9)))))
+ (fun z4 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z4 (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x6 y0) (Zmul (Zmul x5 y1) (Zpos (xO xH)))) (Zmul x4 y2))
+ (Zmul (Zmul x3 y3) (Zpos (xO xH)))) (Zmul x2 y4)) (Zmul (Zmul x1 y5) (Zpos (xO xH))))
+ (Zmul x0 y6))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zmul (Zmul x9 y7) (Zpos (xO xH))) (Zmul x8 y8))
+ (Zmul (Zmul x7 y9) (Zpos (xO xH)))))))
+ (fun z5 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z5 (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x7 y0) (Zmul x6 y1)) (Zmul x5 y2)) (Zmul x4 y3))
+ (Zmul x3 y4)) (Zmul x2 y5)) (Zmul x1 y6)) (Zmul x0 y7))
+ (Zmul (Zpos (xI (xI (xO (xO xH))))) (Zadd (Zmul x9 y8) (Zmul x8 y9)))))
+ (fun z6 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z6 (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x8 y0) (Zmul (Zmul x7 y1) (Zpos (xO xH)))) (Zmul x6 y2))
+ (Zmul (Zmul x5 y3) (Zpos (xO xH)))) (Zmul x4 y4))
+ (Zmul (Zmul x3 y5) (Zpos (xO xH)))) (Zmul x2 y6))
+ (Zmul (Zmul x1 y7) (Zpos (xO xH)))) (Zmul x0 y8))
+ (Zmul (Zmul (Zmul (Zpos (xI (xI (xO (xO xH))))) x9) y9) (Zpos (xO xH)))))
+ (fun z7 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Z.shiftr z7 (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x9 y0) (Zmul x8 y1)) (Zmul x7 y2)) (Zmul x6 y3))
+ (Zmul x5 y4)) (Zmul x4 y5)) (Zmul x3 y6)) (Zmul x2 y7))
+ (Zmul x1 y8)) (Zmul x0 y9)))
+ (fun z8 : Z =>
+ @Let_In Z (fun _ : Z => list Z)
+ (Zadd (Zmul (Zpos (xI (xI (xO (xO xH))))) (Z.shiftr z8 (Zpos (xI (xO (xO (xI xH)))))))
+ (Z.land z
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))))
+ (fun z9 : Z =>
+ @cons Z
+ (Z.land z9
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Zadd (Z.shiftr z9 (Zpos (xO (xI (xO (xI xH))))))
+ (Z.land z0
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z1
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z2
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z3
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z4
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z5
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z6
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z7
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land z8
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@nil Z)))))))))))))))))))))))
+ (@decode
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul
+ (Zmul x9 y1)
+ (Zpos (xO xH)))
+ (Zmul x8 y2))
+ (Zmul
+ (Zmul x7 y3)
+ (Zpos (xO xH))))
+ (Zmul x6 y4))
+ (Zmul (Zmul x5 y5) (Zpos (xO xH))))
+ (Zmul x4 y6))
+ (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8))
+ (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x9 y2) (Zmul x8 y3))
+ (Zmul x7 y4))
+ (Zmul x6 y5))
+ (Zmul x5 y6))
+ (Zmul x4 y7)) (Zmul x3 y8))
+ (Zmul x2 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH))))
+ (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH)))
+ (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH))))
+ (Zmul x6 y6))
+ (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8)) (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2))
+ (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5)) (Zmul x7 y6))
+ (Zmul x6 y7)) (Zmul x5 y8))
+ (Zmul x4 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH))))
+ (Zmul x2 y2)) (Zmul (Zmul x1 y3) (Zpos (xO xH))))
+ (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH))) (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH))))
+ (Zmul x6 y8)) (Zmul (Zmul x5 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2)) (Zmul x2 y3))
+ (Zmul x1 y4)) (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8)) (Zmul x6 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x6 y0) (Zmul (Zmul x5 y1) (Zpos (xO xH)))) (Zmul x4 y2))
+ (Zmul (Zmul x3 y3) (Zpos (xO xH)))) (Zmul x2 y4))
+ (Zmul (Zmul x1 y5) (Zpos (xO xH)))) (Zmul x0 y6))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zmul (Zmul x9 y7) (Zpos (xO xH))) (Zmul x8 y8))
+ (Zmul (Zmul x7 y9) (Zpos (xO xH))))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x7 y0) (Zmul x6 y1)) (Zmul x5 y2)) (Zmul x4 y3))
+ (Zmul x3 y4)) (Zmul x2 y5)) (Zmul x1 y6)) (Zmul x0 y7))
+ (Zmul (Zpos (xI (xI (xO (xO xH))))) (Zadd (Zmul x9 y8) (Zmul x8 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x8 y0) (Zmul (Zmul x7 y1) (Zpos (xO xH)))) (Zmul x6 y2))
+ (Zmul (Zmul x5 y3) (Zpos (xO xH)))) (Zmul x4 y4))
+ (Zmul (Zmul x3 y5) (Zpos (xO xH)))) (Zmul x2 y6)) (Zmul (Zmul x1 y7) (Zpos (xO xH))))
+ (Zmul x0 y8)) (Zmul (Zmul (Zmul (Zpos (xI (xI (xO (xO xH))))) x9) y9) (Zpos (xO xH)))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x9 y0) (Zmul x8 y1)) (Zmul x7 y2)) (Zmul x6 y3)) (Zmul x5 y4))
+ (Zmul x4 y5)) (Zmul x3 y6)) (Zmul x2 y7)) (Zmul x1 y8)) (Zmul x0 y9)))
+ (Zpos (xI (xO (xO (xI xH)))))))
+ (Z.land
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH))) (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH)))) (Zmul x6 y4)) (Zmul (Zmul x5 y5) (Zpos (xO xH))))
+ (Zmul x4 y6)) (Zmul (Zmul x3 y7) (Zpos (xO xH)))) (Zmul x2 y8))
+ (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul
+ (Zmul x9 y1)
+ (Zpos (xO xH)))
+ (Zmul x8 y2))
+ (Zmul
+ (Zmul x7 y3)
+ (Zpos (xO xH))))
+ (Zmul x6 y4))
+ (Zmul
+ (Zmul x5 y5)
+ (Zpos (xO xH))))
+ (Zmul x4 y6))
+ (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8))
+ (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul x9 y2)
+ (Zmul x8 y3))
+ (Zmul x7 y4))
+ (Zmul x6 y5))
+ (Zmul x5 y6))
+ (Zmul x4 y7))
+ (Zmul x3 y8))
+ (Zmul x2 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x2 y0)
+ (Zmul (Zmul x1 y1) (Zpos (xO xH))))
+ (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul (Zmul x9 y3) (Zpos (xO xH)))
+ (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH))))
+ (Zmul x6 y6))
+ (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8))
+ (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2))
+ (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5))
+ (Zmul x7 y6)) (Zmul x6 y7))
+ (Zmul x5 y8)) (Zmul x4 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH))))
+ (Zmul x2 y2)) (Zmul (Zmul x1 y3) (Zpos (xO xH))))
+ (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH))) (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH))))
+ (Zmul x6 y8)) (Zmul (Zmul x5 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2))
+ (Zmul x2 y3)) (Zmul x1 y4)) (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8))
+ (Zmul x6 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x6 y0) (Zmul (Zmul x5 y1) (Zpos (xO xH))))
+ (Zmul x4 y2)) (Zmul (Zmul x3 y3) (Zpos (xO xH))))
+ (Zmul x2 y4)) (Zmul (Zmul x1 y5) (Zpos (xO xH))))
+ (Zmul x0 y6))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zmul (Zmul x9 y7) (Zpos (xO xH))) (Zmul x8 y8))
+ (Zmul (Zmul x7 y9) (Zpos (xO xH))))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x7 y0) (Zmul x6 y1)) (Zmul x5 y2)) (Zmul x4 y3))
+ (Zmul x3 y4)) (Zmul x2 y5)) (Zmul x1 y6)) (Zmul x0 y7))
+ (Zmul (Zpos (xI (xI (xO (xO xH))))) (Zadd (Zmul x9 y8) (Zmul x8 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x8 y0) (Zmul (Zmul x7 y1) (Zpos (xO xH)))) (Zmul x6 y2))
+ (Zmul (Zmul x5 y3) (Zpos (xO xH)))) (Zmul x4 y4))
+ (Zmul (Zmul x3 y5) (Zpos (xO xH)))) (Zmul x2 y6))
+ (Zmul (Zmul x1 y7) (Zpos (xO xH)))) (Zmul x0 y8))
+ (Zmul (Zmul (Zmul (Zpos (xI (xI (xO (xO xH))))) x9) y9) (Zpos (xO xH)))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x9 y0) (Zmul x8 y1)) (Zmul x7 y2)) (Zmul x6 y3))
+ (Zmul x5 y4)) (Zmul x4 y5)) (Zmul x3 y6)) (Zmul x2 y7))
+ (Zmul x1 y8)) (Zmul x0 y9))) (Zpos (xI (xO (xO (xI xH)))))))
+ (Z.land
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH))) (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH)))) (Zmul x6 y4))
+ (Zmul (Zmul x5 y5) (Zpos (xO xH)))) (Zmul x4 y6)) (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH))) (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH)))) (Zmul x6 y4))
+ (Zmul (Zmul x5 y5) (Zpos (xO xH)))) (Zmul x4 y6)) (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH)))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x9 y2) (Zmul x8 y3)) (Zmul x7 y4)) (Zmul x6 y5)) (Zmul x5 y6))
+ (Zmul x4 y7)) (Zmul x3 y8)) (Zmul x2 y9)))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH))) (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH)))) (Zmul x6 y4))
+ (Zmul (Zmul x5 y5) (Zpos (xO xH)))) (Zmul x4 y6))
+ (Zmul (Zmul x3 y7) (Zpos (xO xH)))) (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x9 y2) (Zmul x8 y3)) (Zmul x7 y4)) (Zmul x6 y5))
+ (Zmul x5 y6)) (Zmul x4 y7)) (Zmul x3 y8)) (Zmul x2 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH)))) (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH))) (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH)))) (Zmul x6 y6)) (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8)) (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH))) (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH))))
+ (Zmul x6 y4)) (Zmul (Zmul x5 y5) (Zpos (xO xH))))
+ (Zmul x4 y6)) (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x9 y2) (Zmul x8 y3)) (Zmul x7 y4)) (Zmul x6 y5))
+ (Zmul x5 y6)) (Zmul x4 y7)) (Zmul x3 y8)) (Zmul x2 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH)))) (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH))) (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH)))) (Zmul x6 y6))
+ (Zmul (Zmul x5 y7) (Zpos (xO xH)))) (Zmul x4 y8)) (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2)) (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5)) (Zmul x7 y6)) (Zmul x6 y7)) (Zmul x5 y8))
+ (Zmul x4 y9)))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH))) (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH))))
+ (Zmul x6 y4)) (Zmul (Zmul x5 y5) (Zpos (xO xH))))
+ (Zmul x4 y6)) (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x9 y2) (Zmul x8 y3)) (Zmul x7 y4))
+ (Zmul x6 y5)) (Zmul x5 y6)) (Zmul x4 y7))
+ (Zmul x3 y8)) (Zmul x2 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH)))) (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH))) (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH)))) (Zmul x6 y6))
+ (Zmul (Zmul x5 y7) (Zpos (xO xH)))) (Zmul x4 y8))
+ (Zmul (Zmul x3 y9) (Zpos (xO xH))))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2)) (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5)) (Zmul x7 y6)) (Zmul x6 y7))
+ (Zmul x5 y8)) (Zmul x4 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH)))) (Zmul x2 y2))
+ (Zmul (Zmul x1 y3) (Zpos (xO xH)))) (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH))) (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH)))) (Zmul x6 y8)) (Zmul (Zmul x5 y9) (Zpos (xO xH)))))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y1) (Zpos (xO xH)))
+ (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH))))
+ (Zmul x6 y4)) (Zmul (Zmul x5 y5) (Zpos (xO xH))))
+ (Zmul x4 y6)) (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x9 y2) (Zmul x8 y3)) (Zmul x7 y4))
+ (Zmul x6 y5)) (Zmul x5 y6))
+ (Zmul x4 y7)) (Zmul x3 y8)) (Zmul x2 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH)))) (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH))) (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH))))
+ (Zmul x6 y6)) (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8)) (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2)) (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5)) (Zmul x7 y6)) (Zmul x6 y7))
+ (Zmul x5 y8)) (Zmul x4 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH)))) (Zmul x2 y2))
+ (Zmul (Zmul x1 y3) (Zpos (xO xH)))) (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH))) (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH)))) (Zmul x6 y8))
+ (Zmul (Zmul x5 y9) (Zpos (xO xH))))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2)) (Zmul x2 y3)) (Zmul x1 y4))
+ (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8)) (Zmul x6 y9)))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul (Zmul x9 y1) (Zpos (xO xH)))
+ (Zmul x8 y2))
+ (Zmul (Zmul x7 y3) (Zpos (xO xH))))
+ (Zmul x6 y4))
+ (Zmul (Zmul x5 y5) (Zpos (xO xH))))
+ (Zmul x4 y6))
+ (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8)) (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x9 y2) (Zmul x8 y3))
+ (Zmul x7 y4)) (Zmul x6 y5))
+ (Zmul x5 y6)) (Zmul x4 y7))
+ (Zmul x3 y8)) (Zmul x2 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH))))
+ (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH)))
+ (Zmul x8 y4)) (Zmul (Zmul x7 y5) (Zpos (xO xH))))
+ (Zmul x6 y6)) (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8)) (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2)) (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5)) (Zmul x7 y6))
+ (Zmul x6 y7)) (Zmul x5 y8)) (Zmul x4 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH)))) (Zmul x2 y2))
+ (Zmul (Zmul x1 y3) (Zpos (xO xH)))) (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH))) (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH)))) (Zmul x6 y8))
+ (Zmul (Zmul x5 y9) (Zpos (xO xH))))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2)) (Zmul x2 y3))
+ (Zmul x1 y4)) (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8)) (Zmul x6 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x6 y0) (Zmul (Zmul x5 y1) (Zpos (xO xH)))) (Zmul x4 y2))
+ (Zmul (Zmul x3 y3) (Zpos (xO xH)))) (Zmul x2 y4))
+ (Zmul (Zmul x1 y5) (Zpos (xO xH)))) (Zmul x0 y6))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zmul (Zmul x9 y7) (Zpos (xO xH))) (Zmul x8 y8))
+ (Zmul (Zmul x7 y9) (Zpos (xO xH)))))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul
+ (Zmul x9 y1)
+ (Zpos (xO xH)))
+ (Zmul x8 y2))
+ (Zmul
+ (Zmul x7 y3)
+ (Zpos (xO xH))))
+ (Zmul x6 y4))
+ (Zmul (Zmul x5 y5) (Zpos (xO xH))))
+ (Zmul x4 y6))
+ (Zmul (Zmul x3 y7) (Zpos (xO xH))))
+ (Zmul x2 y8))
+ (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x9 y2) (Zmul x8 y3))
+ (Zmul x7 y4))
+ (Zmul x6 y5))
+ (Zmul x5 y6))
+ (Zmul x4 y7)) (Zmul x3 y8))
+ (Zmul x2 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zmul x2 y0) (Zmul (Zmul x1 y1) (Zpos (xO xH))))
+ (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y3) (Zpos (xO xH)))
+ (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH))))
+ (Zmul x6 y6))
+ (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8)) (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2))
+ (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5)) (Zmul x7 y6))
+ (Zmul x6 y7)) (Zmul x5 y8))
+ (Zmul x4 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH))))
+ (Zmul x2 y2)) (Zmul (Zmul x1 y3) (Zpos (xO xH))))
+ (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH))) (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH))))
+ (Zmul x6 y8)) (Zmul (Zmul x5 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2)) (Zmul x2 y3))
+ (Zmul x1 y4)) (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8)) (Zmul x6 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x6 y0) (Zmul (Zmul x5 y1) (Zpos (xO xH)))) (Zmul x4 y2))
+ (Zmul (Zmul x3 y3) (Zpos (xO xH)))) (Zmul x2 y4))
+ (Zmul (Zmul x1 y5) (Zpos (xO xH)))) (Zmul x0 y6))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zmul (Zmul x9 y7) (Zpos (xO xH))) (Zmul x8 y8))
+ (Zmul (Zmul x7 y9) (Zpos (xO xH))))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zadd (Zmul x7 y0) (Zmul x6 y1)) (Zmul x5 y2)) (Zmul x4 y3))
+ (Zmul x3 y4)) (Zmul x2 y5)) (Zmul x1 y6)) (Zmul x0 y7))
+ (Zmul (Zpos (xI (xI (xO (xO xH))))) (Zadd (Zmul x9 y8) (Zmul x8 y9)))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul
+ (Zmul x9 y1)
+ (Zpos (xO xH)))
+ (Zmul x8 y2))
+ (Zmul
+ (Zmul x7 y3)
+ (Zpos (xO xH))))
+ (Zmul x6 y4))
+ (Zmul
+ (Zmul x5 y5)
+ (Zpos (xO xH))))
+ (Zmul x4 y6))
+ (Zmul
+ (Zmul x3 y7)
+ (Zpos (xO xH))))
+ (Zmul x2 y8))
+ (Zmul (Zmul x1 y9) (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul x9 y2)
+ (Zmul x8 y3))
+ (Zmul x7 y4))
+ (Zmul x6 y5))
+ (Zmul x5 y6))
+ (Zmul x4 y7))
+ (Zmul x3 y8))
+ (Zmul x2 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x2 y0)
+ (Zmul (Zmul x1 y1) (Zpos (xO xH))))
+ (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul
+ (Zmul x9 y3)
+ (Zpos (xO xH)))
+ (Zmul x8 y4))
+ (Zmul (Zmul x7 y5) (Zpos (xO xH))))
+ (Zmul x6 y6))
+ (Zmul (Zmul x5 y7) (Zpos (xO xH))))
+ (Zmul x4 y8))
+ (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1)) (Zmul x1 y2))
+ (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x9 y4) (Zmul x8 y5))
+ (Zmul x7 y6)) (Zmul x6 y7))
+ (Zmul x5 y8)) (Zmul x4 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x4 y0) (Zmul (Zmul x3 y1) (Zpos (xO xH))))
+ (Zmul x2 y2)) (Zmul (Zmul x1 y3) (Zpos (xO xH))))
+ (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH)))
+ (Zmul x8 y6)) (Zmul (Zmul x7 y7) (Zpos (xO xH))))
+ (Zmul x6 y8)) (Zmul (Zmul x5 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2))
+ (Zmul x2 y3)) (Zmul x1 y4)) (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8))
+ (Zmul x6 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x6 y0) (Zmul (Zmul x5 y1) (Zpos (xO xH))))
+ (Zmul x4 y2)) (Zmul (Zmul x3 y3) (Zpos (xO xH))))
+ (Zmul x2 y4)) (Zmul (Zmul x1 y5) (Zpos (xO xH))))
+ (Zmul x0 y6))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zmul (Zmul x9 y7) (Zpos (xO xH))) (Zmul x8 y8))
+ (Zmul (Zmul x7 y9) (Zpos (xO xH))))))) (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x7 y0) (Zmul x6 y1)) (Zmul x5 y2))
+ (Zmul x4 y3)) (Zmul x3 y4)) (Zmul x2 y5))
+ (Zmul x1 y6)) (Zmul x0 y7))
+ (Zmul (Zpos (xI (xI (xO (xO xH))))) (Zadd (Zmul x9 y8) (Zmul x8 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x8 y0) (Zmul (Zmul x7 y1) (Zpos (xO xH))))
+ (Zmul x6 y2)) (Zmul (Zmul x5 y3) (Zpos (xO xH))))
+ (Zmul x4 y4)) (Zmul (Zmul x3 y5) (Zpos (xO xH))))
+ (Zmul x2 y6)) (Zmul (Zmul x1 y7) (Zpos (xO xH)))) (Zmul x0 y8))
+ (Zmul (Zmul (Zmul (Zpos (xI (xI (xO (xO xH))))) x9) y9) (Zpos (xO xH)))))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI xH)))))))))))))))))))))))))))
+ (@cons Z
+ (Z.land
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd
+ (Z.shiftr
+ (Zadd (Zmul x0 y0)
+ (Zmul
+ (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul
+ (Zmul x9 y1)
+ (Zpos (xO xH)))
+ (Zmul x8 y2))
+ (Zmul
+ (Zmul x7 y3)
+ (Zpos (xO xH))))
+ (Zmul x6 y4))
+ (Zmul
+ (Zmul x5 y5)
+ (Zpos (xO xH))))
+ (Zmul x4 y6))
+ (Zmul
+ (Zmul x3 y7)
+ (Zpos (xO xH))))
+ (Zmul x2 y8))
+ (Zmul
+ (Zmul x1 y9)
+ (Zpos (xO xH))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd (Zadd (Zmul x1 y0) (Zmul x0 y1))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul x9 y2)
+ (Zmul x8 y3))
+ (Zmul x7 y4))
+ (Zmul x6 y5))
+ (Zmul x5 y6))
+ (Zmul x4 y7))
+ (Zmul x3 y8))
+ (Zmul x2 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x2 y0)
+ (Zmul (Zmul x1 y1) (Zpos (xO xH))))
+ (Zmul x0 y2))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zmul
+ (Zmul x9 y3)
+ (Zpos (xO xH)))
+ (Zmul x8 y4))
+ (Zmul
+ (Zmul x7 y5)
+ (Zpos (xO xH))))
+ (Zmul x6 y6))
+ (Zmul
+ (Zmul x5 y7)
+ (Zpos (xO xH))))
+ (Zmul x4 y8))
+ (Zmul (Zmul x3 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x3 y0) (Zmul x2 y1))
+ (Zmul x1 y2)) (Zmul x0 y3))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x9 y4) (Zmul x8 y5))
+ (Zmul x7 y6))
+ (Zmul x6 y7))
+ (Zmul x5 y8))
+ (Zmul x4 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x4 y0)
+ (Zmul (Zmul x3 y1) (Zpos (xO xH))))
+ (Zmul x2 y2))
+ (Zmul (Zmul x1 y3) (Zpos (xO xH))))
+ (Zmul x0 y4))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul (Zmul x9 y5) (Zpos (xO xH)))
+ (Zmul x8 y6))
+ (Zmul (Zmul x7 y7) (Zpos (xO xH))))
+ (Zmul x6 y8))
+ (Zmul (Zmul x5 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x5 y0) (Zmul x4 y1)) (Zmul x3 y2))
+ (Zmul x2 y3)) (Zmul x1 y4))
+ (Zmul x0 y5))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zadd (Zmul x9 y6) (Zmul x8 y7)) (Zmul x7 y8))
+ (Zmul x6 y9))))) (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zmul x6 y0)
+ (Zmul (Zmul x5 y1) (Zpos (xO xH))))
+ (Zmul x4 y2)) (Zmul (Zmul x3 y3) (Zpos (xO xH))))
+ (Zmul x2 y4)) (Zmul (Zmul x1 y5) (Zpos (xO xH))))
+ (Zmul x0 y6))
+ (Zmul (Zpos (xI (xI (xO (xO xH)))))
+ (Zadd (Zadd (Zmul (Zmul x9 y7) (Zpos (xO xH))) (Zmul x8 y8))
+ (Zmul (Zmul x7 y9) (Zpos (xO xH)))))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x7 y0) (Zmul x6 y1)) (Zmul x5 y2))
+ (Zmul x4 y3)) (Zmul x3 y4)) (Zmul x2 y5))
+ (Zmul x1 y6)) (Zmul x0 y7))
+ (Zmul (Zpos (xI (xI (xO (xO xH))))) (Zadd (Zmul x9 y8) (Zmul x8 y9)))))
+ (Zpos (xI (xO (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zmul x8 y0) (Zmul (Zmul x7 y1) (Zpos (xO xH))))
+ (Zmul x6 y2)) (Zmul (Zmul x5 y3) (Zpos (xO xH))))
+ (Zmul x4 y4)) (Zmul (Zmul x3 y5) (Zpos (xO xH))))
+ (Zmul x2 y6)) (Zmul (Zmul x1 y7) (Zpos (xO xH))))
+ (Zmul x0 y8))
+ (Zmul (Zmul (Zmul (Zpos (xI (xI (xO (xO xH))))) x9) y9) (Zpos (xO xH)))))
+ (Zpos (xO (xI (xO (xI xH))))))
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd
+ (Zadd (Zadd (Zadd (Zmul x9 y0) (Zmul x8 y1)) (Zmul x7 y2))
+ (Zmul x6 y3)) (Zmul x5 y4)) (Zmul x4 y5))
+ (Zmul x3 y6)) (Zmul x2 y7)) (Zmul x1 y8)) (Zmul x0 y9)))
+ (Zpos
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI
+ (xI (xI (xI (xI (xI (xI (xI xH))))))))))))))))))))))))))
+ (@nil Z)))))))))))).
+ cbv beta zeta.
+ intros.
+ (timeout 1 (apply f_equal; reflexivity)) || fail 0 "too early".
+ Undo.
+ Time Timeout 1 f_equal. (* Finished transaction in 0. secs (0.3u,0.s) in 8.4 *)
diff --git a/test-suite/bugs/closed/4816.v b/test-suite/bugs/closed/4816.v
index 5ba0787ee..00a523842 100644
--- a/test-suite/bugs/closed/4816.v
+++ b/test-suite/bugs/closed/4816.v
@@ -1,11 +1,29 @@
+Section foo.
+Polymorphic Universes A B.
+Fail Constraint A <= B.
+End foo.
+(* gives an anomaly Universe undefined *)
+
+Universes X Y.
+Section Foo.
+ Polymorphic Universes Z W.
+ Polymorphic Constraint W < Z.
+
+ Fail Definition bla := Type@{W}.
+ Polymorphic Definition bla := Type@{W}.
+ Section Bar.
+ Fail Constraint X <= Z.
+ End Bar.
+End Foo.
+
Require Coq.Classes.RelationClasses.
Class PreOrder (A : Type) (r : A -> A -> Type) : Type :=
{ refl : forall x, r x x }.
-Section foo.
+Section qux.
Polymorphic Universes A.
Section bar.
Fail Context {A : Type@{A}} {rA : A -> A -> Prop} {PO : PreOrder A rA}.
End bar.
-End foo. \ No newline at end of file
+End qux.
diff --git a/test-suite/bugs/closed/4865.v b/test-suite/bugs/closed/4865.v
new file mode 100644
index 000000000..c5bf3289b
--- /dev/null
+++ b/test-suite/bugs/closed/4865.v
@@ -0,0 +1,52 @@
+(* Check discharge of arguments scopes + other checks *)
+
+(* This is bug #4865 *)
+
+Notation "<T>" := true : bool_scope.
+Section A.
+ Check negb <T>.
+ Global Arguments negb : clear scopes.
+ Fail Check negb <T>.
+End A.
+
+(* Check that no scope is re-computed *)
+Fail Check negb <T>.
+
+(* Another test about arguments scopes in sections *)
+
+Notation "0" := true.
+Section B.
+ Variable x : nat.
+ Let T := nat -> nat.
+ Definition f y : T := fun z => x + y + z.
+ Fail Check f 1 0. (* 0 in nat, 0 in bool *)
+ Fail Check f 0 0. (* 0 in nat, 0 in bool *)
+ Check f 0 1. (* 0 and 1 in nat *)
+ Global Arguments f _%nat_scope _%nat_scope.
+ Check f 0 0. (* both 0 in nat *)
+End B.
+
+(* Check that only the scope for the extra product on x is re-computed *)
+Check f 0 0 0. (* All 0 in nat *)
+
+Section C.
+ Variable x : nat.
+ Let T := nat -> nat.
+ Definition g y : T := fun z => x + y + z.
+ Global Arguments g : clear scopes.
+ Check g 1. (* 1 in nat *)
+End C.
+
+(* Check that only the scope for the extra product on x is re-computed *)
+Check g 0. (* 0 in nat *)
+Fail Check g 0 1 0. (* 2nd 0 in bool *)
+Fail Check g 0 0 1. (* 2nd 0 in bool *)
+
+(* Another test on arguments scopes: checking scope for expanding arities *)
+(* Not sure this is very useful, but why not *)
+
+Fixpoint arr n := match n with 0%nat => nat | S n => nat -> arr n end.
+Fixpoint lam n : arr n := match n with 0%nat => 0%nat | S n => fun x => lam n end.
+Notation "0" := true.
+Arguments Scope lam [nat_scope nat_scope].
+Check (lam 1 0).
diff --git a/test-suite/bugs/closed/4873.v b/test-suite/bugs/closed/4873.v
new file mode 100644
index 000000000..f2f917b4e
--- /dev/null
+++ b/test-suite/bugs/closed/4873.v
@@ -0,0 +1,72 @@
+Require Import Coq.Classes.Morphisms.
+Require Import Relation_Definitions.
+Require Import Coq.Compat.Coq85.
+
+Fixpoint tuple' T n : Type :=
+ match n with
+ | O => T
+ | S n' => (tuple' T n' * T)%type
+ end.
+
+Definition tuple T n : Type :=
+ match n with
+ | O => unit
+ | S n' => tuple' T n'
+ end.
+
+Fixpoint to_list' {T} (n:nat) {struct n} : tuple' T n -> list T :=
+ match n with
+ | 0 => fun x => (x::nil)%list
+ | S n' => fun xs : tuple' T (S n') => let (xs', x) := xs in (x :: to_list' n' xs')%list
+ end.
+
+Definition to_list {T} (n:nat) : tuple T n -> list T :=
+ match n with
+ | 0 => fun _ => nil
+ | S n' => fun xs : tuple T (S n') => to_list' n' xs
+ end.
+
+Program Fixpoint from_list' {T} (y:T) (n:nat) (xs:list T) : length xs = n -> tuple' T n :=
+ match n return _ with
+ | 0 =>
+ match xs return (length xs = 0 -> tuple' T 0) with
+ | nil => fun _ => y
+ | _ => _ (* impossible *)
+ end
+ | S n' =>
+ match xs return (length xs = S n' -> tuple' T (S n')) with
+ | cons x xs' => fun _ => (from_list' x n' xs' _, y)
+ | _ => _ (* impossible *)
+ end
+ end.
+Goal True.
+ pose from_list'_obligation_3 as e.
+ repeat (let e' := fresh in
+ rename e into e';
+ (pose (e' nat) as e || pose (e' 0) as e || pose (e' nil) as e || pose (e' eq_refl) as e);
+ subst e').
+ progress hnf in e.
+ pose (eq_refl : e = eq_refl).
+ exact I.
+Qed.
+
+Program Definition from_list {T} (n:nat) (xs:list T) : length xs = n -> tuple T n :=
+match n return _ with
+| 0 =>
+ match xs return (length xs = 0 -> tuple T 0) with
+ | nil => fun _ : 0 = 0 => tt
+ | _ => _ (* impossible *)
+ end
+| S n' =>
+ match xs return (length xs = S n' -> tuple T (S n')) with
+ | cons x xs' => fun _ => from_list' x n' xs' _
+ | _ => _ (* impossible *)
+ end
+end.
+
+Lemma to_list_from_list : forall {T} (n:nat) (xs:list T) pf, to_list n (from_list n xs pf) = xs.
+Proof.
+ destruct xs; simpl; intros; subst; auto.
+ generalize dependent t. simpl in *.
+ induction xs; simpl in *; intros; congruence.
+Qed. \ No newline at end of file
diff --git a/test-suite/bugs/closed/HoTT_coq_047.v b/test-suite/bugs/closed/HoTT_coq_047.v
index 29496be5e..bef3c33ca 100644
--- a/test-suite/bugs/closed/HoTT_coq_047.v
+++ b/test-suite/bugs/closed/HoTT_coq_047.v
@@ -1,3 +1,5 @@
+Unset Structural Injection.
+
Inductive nCk : nat -> nat -> Type :=
|zz : nCk 0 0
| incl { m n : nat } : nCk m n -> nCk (S m) (S n)
diff --git a/test-suite/bugs/opened/3410.v b/test-suite/bugs/opened/3410.v
deleted file mode 100644
index 0d259181a..000000000
--- a/test-suite/bugs/opened/3410.v
+++ /dev/null
@@ -1 +0,0 @@
-Fail repeat match goal with H:_ |- _ => setoid_rewrite X in H end.
diff --git a/test-suite/bugs/opened/3889.v b/test-suite/bugs/opened/3889.v
new file mode 100644
index 000000000..6b287324c
--- /dev/null
+++ b/test-suite/bugs/opened/3889.v
@@ -0,0 +1,11 @@
+Require Import Program.
+
+Inductive Even : nat -> Prop :=
+| evenO : Even O
+| evenS : forall n, Odd n -> Even (S n)
+with Odd : nat -> Prop :=
+| oddS : forall n, Even n -> Odd (S n).
+Axiom admit : forall {T}, T.
+Program Fixpoint doubleE {n} (e : Even n) : Even (2 * n) := admit
+with doubleO {n} (o : Odd n) : Odd (S (2 * n)) := _.
+Next Obligation of doubleE.
diff --git a/test-suite/bugs/opened/3890.v b/test-suite/bugs/opened/3890.v
new file mode 100644
index 000000000..f9ac9be2c
--- /dev/null
+++ b/test-suite/bugs/opened/3890.v
@@ -0,0 +1,18 @@
+Class Foo.
+Class Bar := b : Type.
+
+Instance foo : Foo := _.
+(* 1 subgoals, subgoal 1 (ID 4)
+
+ ============================
+ Foo *)
+
+Instance bar : Bar.
+exact Type.
+Defined.
+(* bar is defined *)
+
+About foo.
+(* foo not a defined object. *)
+
+Fail Defined.
diff --git a/test-suite/bugs/opened/3916.v b/test-suite/bugs/opened/3916.v
new file mode 100644
index 000000000..fd95503e6
--- /dev/null
+++ b/test-suite/bugs/opened/3916.v
@@ -0,0 +1,3 @@
+Require Import List.
+
+Fail Hint Resolve -> in_map. (* Also happens when using <- instead of -> *)
diff --git a/test-suite/bugs/opened/3919.v-disabled b/test-suite/bugs/opened/3919.v-disabled
new file mode 100644
index 000000000..0d661de9c
--- /dev/null
+++ b/test-suite/bugs/opened/3919.v-disabled
@@ -0,0 +1,13 @@
+Require Import MSets.
+Require Import Orders.
+
+Declare Module Signal : OrderedType.
+
+Module S := MSetAVL.Make(Signal).
+Module Sdec := Decide(S).
+Export Sdec.
+
+Hint Extern 0 (Signal.eq ?x ?y) => now symmetry.
+
+Goal forall o s, Signal.eq o s.
+Proof. fsetdec. Qed.
diff --git a/test-suite/bugs/opened/3922.v-disabled b/test-suite/bugs/opened/3922.v-disabled
new file mode 100644
index 000000000..ce4f509ca
--- /dev/null
+++ b/test-suite/bugs/opened/3922.v-disabled
@@ -0,0 +1,83 @@
+Set Universe Polymorphism.
+Notation Type0 := Set.
+
+Definition Type1 := Eval hnf in let gt := (Set : Type@{i}) in Type@{i}.
+
+Notation compose := (fun g f x => g (f x)).
+
+Notation "g 'o' f" := (compose g f) (at level 40, left associativity) : function_scope.
+Open Scope function_scope.
+
+Definition pointwise_paths {A} {P:A->Type} (f g:forall x:A, P x)
+ := forall x:A, f x = g x.
+
+Notation "f == g" := (pointwise_paths f g) (at level 70, no associativity) : type_scope.
+
+Class Contr_internal (A : Type) := BuildContr {
+ center : A ;
+ contr : (forall y : A, center = y)
+}.
+
+Inductive trunc_index : Type :=
+| minus_two : trunc_index
+| trunc_S : trunc_index -> trunc_index.
+
+Notation "n .+1" := (trunc_S n) (at level 2, left associativity, format "n .+1") : trunc_scope.
+Local Open Scope trunc_scope.
+Notation "-2" := minus_two (at level 0) : trunc_scope.
+Notation "-1" := (-2.+1) (at level 0) : trunc_scope.
+
+Fixpoint IsTrunc_internal (n : trunc_index) (A : Type) : Type :=
+ match n with
+ | -2 => Contr_internal A
+ | n'.+1 => forall (x y : A), IsTrunc_internal n' (x = y)
+ end.
+
+Class IsTrunc (n : trunc_index) (A : Type) : Type :=
+ Trunc_is_trunc : IsTrunc_internal n A.
+
+Notation Contr := (IsTrunc -2).
+Notation IsHProp := (IsTrunc -1).
+
+Monomorphic Axiom dummy_funext_type : Type0.
+Monomorphic Class Funext := { dummy_funext_value : dummy_funext_type }.
+
+Inductive Unit : Type1 :=
+ tt : Unit.
+
+Record TruncType (n : trunc_index) := BuildTruncType {
+ trunctype_type : Type ;
+ istrunc_trunctype_type : IsTrunc n trunctype_type
+}.
+
+Arguments BuildTruncType _ _ {_}.
+
+Coercion trunctype_type : TruncType >-> Sortclass.
+
+Notation "n -Type" := (TruncType n) (at level 1) : type_scope.
+Notation hProp := (-1)-Type.
+
+Notation BuildhProp := (BuildTruncType -1).
+
+Private Inductive Trunc (n : trunc_index) (A :Type) : Type :=
+ tr : A -> Trunc n A.
+Arguments tr {n A} a.
+
+Global Instance istrunc_truncation (n : trunc_index) (A : Type@{i})
+: IsTrunc@{j} n (Trunc@{i} n A).
+Admitted.
+
+Definition Trunc_ind {n A}
+ (P : Trunc n A -> Type) {Pt : forall aa, IsTrunc n (P aa)}
+ : (forall a, P (tr a)) -> (forall aa, P aa)
+:= (fun f aa => match aa with tr a => fun _ => f a end Pt).
+Definition merely (A : Type@{i}) : hProp@{i} := BuildhProp (Trunc -1 A).
+Definition cconst_factors_contr `{Funext} {X Y : Type} (f : X -> Y)
+ (P : Type) `{Pc : X -> Contr P}
+ (g : X -> P) (h : P -> Y) (p : h o g == f)
+: Unit.
+Proof.
+ assert (merely X -> IsHProp P) by admit.
+ refine (let g' := Trunc_ind (fun _ => P) g : merely X -> P in _);
+ [ assumption.. | ].
+ Fail pose (g' := Trunc_ind (fun _ => P) g : merely X -> P).
diff --git a/test-suite/bugs/opened/3926.v b/test-suite/bugs/opened/3926.v
new file mode 100644
index 000000000..cfad76357
--- /dev/null
+++ b/test-suite/bugs/opened/3926.v
@@ -0,0 +1,30 @@
+Notation compose := (fun g f x => g (f x)).
+Notation "g 'o' f" := (compose g f) (at level 40, left associativity) : function_scope.
+Open Scope function_scope.
+Inductive paths {A : Type} (a : A) : A -> Type := idpath : paths a a where "x = y" := (@paths _ x y) : type_scope.
+Arguments idpath {A a} , [A] a.
+Definition ap {A B:Type} (f:A -> B) {x y:A} (p:x = y) : f x = f y := match p with idpath => idpath end.
+Class IsEquiv {A B : Type} (f : A -> B) := { equiv_inv : B -> A }.
+Notation "f ^-1" := (@equiv_inv _ _ f _) (at level 3, format "f '^-1'") : equiv_scope.
+Local Open Scope equiv_scope.
+Axiom eisretr : forall {A B} (f : A -> B) `{IsEquiv A B f} x, f (f^-1 x) = x.
+Generalizable Variables A B C f g.
+Global Instance isequiv_compose `{IsEquiv A B f} `{IsEquiv B C g} : IsEquiv (compose g f) | 1000
+ := Build_IsEquiv A C (compose g f) (compose f^-1 g^-1).
+Definition isequiv_homotopic {A B} (f : A -> B) {g : A -> B} `{IsEquiv A B f} (h : forall x, f x = g x) : IsEquiv g
+ := Build_IsEquiv _ _ g (f ^-1).
+Global Instance isequiv_inverse {A B} (f : A -> B) `{IsEquiv A B f} : IsEquiv f^-1 | 10000
+ := Build_IsEquiv B A f^-1 f.
+Definition cancelR_isequiv {A B C} (f : A -> B) {g : B -> C}
+ `{IsEquiv A B f} `{IsEquiv A C (g o f)}
+ : IsEquiv g.
+Proof.
+ Unset Typeclasses Modulo Eta.
+ exact (isequiv_homotopic (compose (compose g f) f^-1)
+ (fun b => ap g (eisretr f b))) || fail "too early".
+ Undo.
+ Set Typeclasses Modulo Eta.
+ Set Typeclasses Dependency Order.
+ Set Typeclasses Debug.
+ Fail exact (isequiv_homotopic (compose (compose g f) f^-1)
+ (fun b => ap g (eisretr f b))).
diff --git a/test-suite/bugs/opened/3928.v-disabled b/test-suite/bugs/opened/3928.v-disabled
new file mode 100644
index 000000000..b470eb229
--- /dev/null
+++ b/test-suite/bugs/opened/3928.v-disabled
@@ -0,0 +1,12 @@
+Typeclasses eauto := bfs.
+
+Class Foo := {}.
+Class Bar := {}.
+
+Instance: Bar.
+Instance: Foo -> Bar -> Foo -> Foo | 1.
+Instance: Bar -> Foo | 100.
+Instance: Foo -> Bar -> Foo -> Foo | 1.
+
+Set Typeclasses Debug.
+Timeout 1 Check (_ : Foo). (* timeout *)
diff --git a/test-suite/bugs/opened/3938.v b/test-suite/bugs/opened/3938.v
new file mode 100644
index 000000000..2d0d1930f
--- /dev/null
+++ b/test-suite/bugs/opened/3938.v
@@ -0,0 +1,6 @@
+Require Import Coq.Arith.PeanoNat.
+Hint Extern 1 => admit : typeclass_instances.
+Goal forall a b (f : nat -> Set), Nat.eq a b -> f a = f b.
+ intros a b f H.
+ rewrite H. (* Toplevel input, characters 15-25:
+Anomaly: Evar ?X11 was not declared. Please report. *)
diff --git a/test-suite/bugs/opened/3946.v b/test-suite/bugs/opened/3946.v
new file mode 100644
index 000000000..e77bdbc65
--- /dev/null
+++ b/test-suite/bugs/opened/3946.v
@@ -0,0 +1,11 @@
+Require Import ZArith.
+
+Inductive foo := Foo : Z.le 0 1 -> foo.
+
+Definition bar (f : foo) := let (f) := f in f.
+
+(* Doesn't work: *)
+(* Arguments bar f.*)
+
+(* Does work *)
+Arguments bar f _.
diff --git a/test-suite/bugs/opened/3948.v b/test-suite/bugs/opened/3948.v
new file mode 100644
index 000000000..165813084
--- /dev/null
+++ b/test-suite/bugs/opened/3948.v
@@ -0,0 +1,25 @@
+Module Type S.
+Parameter t : Type.
+End S.
+
+Module Bar(X : S).
+Proof.
+ Definition elt := X.t.
+ Axiom fold : elt.
+End Bar.
+
+Module Make (X: S) := Bar(X).
+
+Declare Module X : S.
+
+Module Type Interface.
+ Parameter constant : unit.
+End Interface.
+
+Module DepMap : Interface.
+ Module Dom := Make(X).
+ Definition constant : unit :=
+ let _ := @Dom.fold in tt.
+End DepMap.
+
+Print Assumptions DepMap.constant. \ No newline at end of file
diff --git a/test-suite/bugs/opened/4813.v b/test-suite/bugs/opened/4813.v
new file mode 100644
index 000000000..b75170179
--- /dev/null
+++ b/test-suite/bugs/opened/4813.v
@@ -0,0 +1,5 @@
+(* An example one would like to see succeeding *)
+
+Record T := BT { t : Set }.
+Record U (x : T) := BU { u : t x -> Prop }.
+Fail Definition A (H : unit -> Prop) : U (BT unit) := BU _ H.